src/HOL/Analysis/Conformal_Mappings.thy
author haftmann
Fri, 14 Jun 2019 08:34:28 +0000
changeset 70347 e5cd5471c18a
parent 70346 408e15cbd2a6
child 70365 4df0628e8545
permissions -rw-r--r--
official fact collection sign_simps
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
69517
dc20f278e8f3 tuned style and headers
nipkow
parents: 69516
diff changeset
     1
section \<open>Conformal Mappings and Consequences of Cauchy's Integral Theorem\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2016)\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
     5
text\<open>Also Cauchy's residue theorem by Wenda Li (2016)\<close>
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
     6
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
theory Conformal_Mappings
65538
a39ef48fbee0 tuned imports;
wenzelm
parents: 65064
diff changeset
     8
imports Cauchy_Integral_Theorem
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
    12
(* FIXME mv to Cauchy_Integral_Theorem.thy *)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
subsection\<open>Cauchy's inequality and more versions of Liouville\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
lemma Cauchy_higher_deriv_bound:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
    assumes holf: "f holomorphic_on (ball z r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
        and contf: "continuous_on (cball z r) f"
65036
ab7e11730ad8 Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
    18
        and fin : "\<And>w. w \<in> ball z r \<Longrightarrow> f w \<in> ball y B0"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
        and "0 < r" and "0 < n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
      shows "norm ((deriv ^^ n) f z) \<le> (fact n) * B0 / r^n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
  have "0 < B0" using \<open>0 < r\<close> fin [of z]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
    by (metis ball_eq_empty ex_in_conv fin not_less)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  have le_B0: "\<And>w. cmod (w - z) \<le> r \<Longrightarrow> cmod (f w - y) \<le> B0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
    apply (rule continuous_on_closure_norm_le [of "ball z r" "\<lambda>w. f w - y"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
    apply (auto simp: \<open>0 < r\<close>  dist_norm norm_minus_commute)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
    apply (rule continuous_intros contf)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
    using fin apply (simp add: dist_commute dist_norm less_eq_real_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w) z - (deriv ^^ n) (\<lambda>w. y) z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
    using \<open>0 < n\<close> by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  also have "... = (deriv ^^ n) (\<lambda>w. f w - y) z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
    by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \<open>0 < r\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
  finally have "(deriv ^^ n) f z = (deriv ^^ n) (\<lambda>w. f w - y) z" .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
  have contf': "continuous_on (cball z r) (\<lambda>u. f u - y)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
    by (rule contf continuous_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  have holf': "(\<lambda>u. (f u - y)) holomorphic_on (ball z r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
    by (simp add: holf holomorphic_on_diff)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
    39
  define a where "a = (2 * pi)/(fact n)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
  have "0 < a"  by (simp add: a_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
  have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
    using \<open>0 < r\<close> by (simp add: a_def divide_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
  have der_dif: "(deriv ^^ n) (\<lambda>w. f w - y) z = (deriv ^^ n) f z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
    using \<open>0 < r\<close> \<open>0 < n\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
    by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const])
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
    46
  have "norm ((2 * of_real pi * \<i>)/(fact n) * (deriv ^^ n) (\<lambda>w. f w - y) z)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
        \<le> (B0/r^(Suc n)) * (2 * pi * r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
    apply (rule has_contour_integral_bound_circlepath [of "(\<lambda>u. (f u - y)/(u - z)^(Suc n))" _ z])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
    using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf']
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
    using \<open>0 < B0\<close> \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
    apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
    using \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
    by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
    58
lemma Cauchy_inequality:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
    assumes holf: "f holomorphic_on (ball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
        and contf: "continuous_on (cball \<xi> r) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
        and "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
        and nof: "\<And>x. norm(\<xi>-x) = r \<Longrightarrow> norm(f x) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
      shows "norm ((deriv ^^ n) f \<xi>) \<le> (fact n) * B / r^n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  obtain x where "norm (\<xi>-x) = r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
    by (metis abs_of_nonneg add_diff_cancel_left' \<open>0 < r\<close> diff_add_cancel
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
                 dual_order.strict_implies_order norm_of_real)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  then have "0 \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
    by (metis nof norm_not_less_zero not_le order_trans)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  have  "((\<lambda>u. f u / (u - \<xi>) ^ Suc n) has_contour_integral (2 * pi) * \<i> / fact n * (deriv ^^ n) f \<xi>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
         (circlepath \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
    apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
    using \<open>0 < r\<close> by simp
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
    74
  then have "norm ((2 * pi * \<i>)/(fact n) * (deriv ^^ n) f \<xi>) \<le> (B / r^(Suc n)) * (2 * pi * r)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
    apply (rule has_contour_integral_bound_circlepath)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
    using \<open>0 \<le> B\<close> \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
    apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  then show ?thesis using \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
    by (simp add: norm_divide norm_mult field_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
    83
lemma Liouville_polynomial:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
    assumes holf: "f holomorphic_on UNIV"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
        and nof: "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) \<le> B * norm z ^ n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
      shows "f \<xi> = (\<Sum>k\<le>n. (deriv^^k) f 0 / fact k * \<xi> ^ k)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
proof (cases rule: le_less_linear [THEN disjE])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
  assume "B \<le> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
  then have "\<And>z. A \<le> norm z \<Longrightarrow> norm(f z) = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
    by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  then have f0: "(f \<longlongrightarrow> 0) at_infinity"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
    using Lim_at_infinity by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
  then have [simp]: "f = (\<lambda>w. 0)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
    using Liouville_weak [OF holf, of 0]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    by (simp add: eventually_at_infinity f0) meson
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  show ?thesis by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
  assume "0 < B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  have "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * (\<xi> - 0)^k) sums f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    apply (rule holomorphic_power_series [where r = "norm \<xi> + 1"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
    using holf holomorphic_on_subset apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
  then have sumsf: "((\<lambda>k. (deriv ^^ k) f 0 / (fact k) * \<xi>^k) sums f \<xi>)" by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  have "(deriv ^^ k) f 0 / fact k * \<xi> ^ k = 0" if "k>n" for k
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  proof (cases "(deriv ^^ k) f 0 = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
    case True then show ?thesis by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   109
    define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   110
    have "1 \<le> abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   111
      using \<open>0 < B\<close> by simp
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   112
    then have wge1: "1 \<le> norm w"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   113
      by (metis norm_of_real w_def)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   114
    then have "w \<noteq> 0" by auto
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   115
    have kB: "0 < fact k * B"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   116
      using \<open>0 < B\<close> by simp
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   117
    then have "0 \<le> fact k * B / cmod ((deriv ^^ k) f 0)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   118
      by simp
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   119
    then have wgeA: "A \<le> cmod w"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   120
      by (simp only: w_def norm_of_real)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   121
    have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\<bar>A\<bar> + 1))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   122
      using \<open>0 < B\<close> by simp
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   123
    then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   124
      by (metis norm_of_real w_def)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   125
    then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   126
      using False by (simp add: divide_simps mult.commute split: if_split_asm)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   127
    also have "... \<le> fact k * (B * norm w ^ n) / norm w ^ k"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   128
      apply (rule Cauchy_inequality)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   129
         using holf holomorphic_on_subset apply force
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   130
        using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   131
       using \<open>w \<noteq> 0\<close> apply simp
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   132
       by (metis nof wgeA dist_0_norm dist_norm)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   133
    also have "... = fact k * (B * 1 / cmod w ^ (k-n))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   134
      apply (simp only: mult_cancel_left times_divide_eq_right [symmetric])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   135
      using \<open>k>n\<close> \<open>w \<noteq> 0\<close> \<open>0 < B\<close> apply (simp add: divide_simps semiring_normalization_rules)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   136
      done
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   137
    also have "... = fact k * B / cmod w ^ (k-n)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   138
      by simp
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   139
    finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" .
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   140
    then have "1 / cmod w < 1 / cmod w ^ (k - n)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   141
      by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   142
    then have "cmod w ^ (k - n) < cmod w"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   143
      by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   144
    with self_le_power [OF wge1] have False
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   145
      by (meson diff_is_0_eq not_gr0 not_le that)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   146
    then show ?thesis by blast
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
  then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \<xi> ^ (k + Suc n) = 0" for k
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
    using not_less_eq by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
  then have "(\<lambda>i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \<xi> ^ (i + Suc n)) sums 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
    by (rule sums_0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  with sums_split_initial_segment [OF sumsf, where n = "Suc n"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
    using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
text\<open>Every bounded entire function is a constant function.\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
theorem Liouville_theorem:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
    assumes holf: "f holomorphic_on UNIV"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
        and bf: "bounded (range f)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
    obtains c where "\<And>z. f z = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  obtain B where "\<And>z. cmod (f z) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
    by (meson bf bounded_pos rangeI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
    using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
text\<open>A holomorphic function f has only isolated zeros unless f is 0.\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   171
lemma powser_0_nonzero:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  fixes a :: "nat \<Rightarrow> 'a::{real_normed_field,banach}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  assumes r: "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
      and sm: "\<And>x. norm (x - \<xi>) < r \<Longrightarrow> (\<lambda>n. a n * (x - \<xi>) ^ n) sums (f x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
      and [simp]: "f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
      and m0: "a m \<noteq> 0" and "m>0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  obtains s where "0 < s" and "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  have "r \<le> conv_radius a"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
    using sm sums_summable by (auto simp: le_conv_radius_iff [where \<xi>=\<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  obtain m where am: "a m \<noteq> 0" and az [simp]: "(\<And>n. n<m \<Longrightarrow> a n = 0)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
    apply (rule_tac m = "LEAST n. a n \<noteq> 0" in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
    using m0
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
    apply (rule LeastI2)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    apply (fastforce intro:  dest!: not_less_Least)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
    done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   187
  define b where "b i = a (i+m) / a m" for i
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   188
  define g where "g x = suminf (\<lambda>i. b i * (x - \<xi>) ^ i)" for x
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
  have [simp]: "b 0 = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
    by (simp add: am b_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  { fix x::'a
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    assume "norm (x - \<xi>) < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
    then have "(\<lambda>n. (a m * (x - \<xi>)^m) * (b n * (x - \<xi>)^n)) sums (f x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
      using am az sm sums_zero_iff_shift [of m "(\<lambda>n. a n * (x - \<xi>) ^ n)" "f x"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
      by (simp add: b_def monoid_mult_class.power_add algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
    then have "x \<noteq> \<xi> \<Longrightarrow> (\<lambda>n. b n * (x - \<xi>)^n) sums (f x / (a m * (x - \<xi>)^m))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
      using am by (simp add: sums_mult_D)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  } note bsums = this
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  then have  "norm (x - \<xi>) < r \<Longrightarrow> summable (\<lambda>n. b n * (x - \<xi>)^n)" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
    using sums_summable by (cases "x=\<xi>") auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  then have "r \<le> conv_radius b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
    by (simp add: le_conv_radius_iff [where \<xi>=\<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  then have "r/2 < conv_radius b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    using not_le order_trans r by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  then have "continuous_on (cball \<xi> (r/2)) g"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
    using powser_continuous_suminf [of "r/2" b \<xi>] by (simp add: g_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
  then obtain s where "s>0"  "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> dist (g x) (g \<xi>) < 1/2"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    apply (rule continuous_onE [where x=\<xi> and e = "1/2"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
    using r apply (auto simp: norm_minus_commute dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
  moreover have "g \<xi> = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
    by (simp add: g_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
  ultimately have gnz: "\<And>x. \<lbrakk>norm (x - \<xi>) \<le> s; norm (x - \<xi>) \<le> r/2\<rbrakk> \<Longrightarrow> (g x) \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
    by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  have "f x \<noteq> 0" if "x \<noteq> \<xi>" "norm (x - \<xi>) \<le> s" "norm (x - \<xi>) \<le> r/2" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
    using bsums [of x] that gnz [of x]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
    apply (auto simp: g_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
    using r sums_iff by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
    apply (rule_tac s="min s (r/2)" in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
    using \<open>0 < r\<close> \<open>0 < s\<close> by (auto simp: dist_commute dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   224
subsection \<open>Analytic continuation\<close>
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   225
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
proposition isolated_zeros:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
      and "open S" "connected S" "\<xi> \<in> S" "f \<xi> = 0" "\<beta> \<in> S" "f \<beta> \<noteq> 0"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   229
    obtains r where "0 < r" and "ball \<xi> r \<subseteq> S" and 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   230
        "\<And>z. z \<in> ball \<xi> r - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
  obtain r where "0 < r" and r: "ball \<xi> r \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
    using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_ball_eq by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
  have powf: "((\<lambda>n. (deriv ^^ n) f \<xi> / (fact n) * (z - \<xi>)^n) sums f z)" if "z \<in> ball \<xi> r" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
    apply (rule holomorphic_power_series [OF _ that])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
    apply (rule holomorphic_on_subset [OF holf r])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
  obtain m where m: "(deriv ^^ m) f \<xi> / (fact m) \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
    using holomorphic_fun_eq_0_on_connected [OF holf \<open>open S\<close> \<open>connected S\<close> _ \<open>\<xi> \<in> S\<close> \<open>\<beta> \<in> S\<close>] \<open>f \<beta> \<noteq> 0\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
  then have "m \<noteq> 0" using assms(5) funpow_0 by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
  obtain s where "0 < s" and s: "\<And>z. z \<in> cball \<xi> s - {\<xi>} \<Longrightarrow> f z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
    apply (rule powser_0_nonzero [OF \<open>0 < r\<close> powf \<open>f \<xi> = 0\<close> m])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
    using \<open>m \<noteq> 0\<close> by (auto simp: dist_commute dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  have "0 < min r s"  by (simp add: \<open>0 < r\<close> \<open>0 < s\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
    apply (rule that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
    using r s by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
proposition analytic_continuation:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  assumes holf: "f holomorphic_on S"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   253
      and "open S" and "connected S"
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   254
      and "U \<subseteq> S" and "\<xi> \<in> S"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
      and "\<xi> islimpt U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
      and fU0 [simp]: "\<And>z. z \<in> U \<Longrightarrow> f z = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
      and "w \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
    shows "f w = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
  obtain e where "0 < e" and e: "cball \<xi> e \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
    using \<open>open S\<close> \<open>\<xi> \<in> S\<close> open_contains_cball_eq by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   262
  define T where "T = cball \<xi> e \<inter> U"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  have contf: "continuous_on (closure T) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
    by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
              holomorphic_on_subset inf.cobounded1)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  have fT0 [simp]: "\<And>x. x \<in> T \<Longrightarrow> f x = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    by (simp add: T_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  have "\<And>r. \<lbrakk>\<forall>e>0. \<exists>x'\<in>U. x' \<noteq> \<xi> \<and> dist x' \<xi> < e; 0 < r\<rbrakk> \<Longrightarrow> \<exists>x'\<in>cball \<xi> e \<inter> U. x' \<noteq> \<xi> \<and> dist x' \<xi> < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    by (metis \<open>0 < e\<close> IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
  then have "\<xi> islimpt T" using \<open>\<xi> islimpt U\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    by (auto simp: T_def islimpt_approachable)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
  then have "\<xi> \<in> closure T"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
    by (simp add: closure_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
  then have "f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
    by (auto simp: continuous_constant_on_closure [OF contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
    apply (rule ccontr)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    apply (rule isolated_zeros [OF holf \<open>open S\<close> \<open>connected S\<close> \<open>\<xi> \<in> S\<close> \<open>f \<xi> = 0\<close> \<open>w \<in> S\<close>], assumption)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
    by (metis open_ball \<open>\<xi> islimpt T\<close> centre_in_ball fT0 insertE insert_Diff islimptE)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   282
corollary analytic_continuation_open:
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   283
  assumes "open s" and "open s'" and "s \<noteq> {}" and "connected s'" 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   284
      and "s \<subseteq> s'"
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   285
  assumes "f holomorphic_on s'" and "g holomorphic_on s'" 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   286
      and "\<And>z. z \<in> s \<Longrightarrow> f z = g z"
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   287
  assumes "z \<in> s'"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   288
  shows   "f z = g z"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   289
proof -
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   290
  from \<open>s \<noteq> {}\<close> obtain \<xi> where "\<xi> \<in> s" by auto
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   291
  with \<open>open s\<close> have \<xi>: "\<xi> islimpt s" 
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   292
    by (intro interior_limit_point) (auto simp: interior_open)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   293
  have "f z - g z = 0"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   294
    by (rule analytic_continuation[of "\<lambda>z. f z - g z" s' s \<xi>])
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   295
       (insert assms \<open>\<xi> \<in> s\<close> \<xi>, auto intro: holomorphic_intros)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   296
  thus ?thesis by simp
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   297
qed
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
   298
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
subsection\<open>Open mapping theorem\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
lemma holomorphic_contract_to_zero:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
  assumes contf: "continuous_on (cball \<xi> r) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
      and holf: "f holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
      and "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
      and norm_less: "\<And>z. norm(\<xi> - z) = r \<Longrightarrow> norm(f \<xi>) < norm(f z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
  obtains z where "z \<in> ball \<xi> r" "f z = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
  { assume fnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
    then have "0 < norm (f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
      by (simp add: \<open>0 < r\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
    have fnz': "\<And>w. w \<in> cball \<xi> r \<Longrightarrow> f w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
      by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
    have "frontier(cball \<xi> r) \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
      using \<open>0 < r\<close> by simp
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   315
    define g where [abs_def]: "g z = inverse (f z)" for z
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
    have contg: "continuous_on (cball \<xi> r) g"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
      unfolding g_def using contf continuous_on_inverse fnz' by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
    have holg: "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
      unfolding g_def using fnz holf holomorphic_on_inverse by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
    have "frontier (cball \<xi> r) \<subseteq> cball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
      by (simp add: subset_iff)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
    then have contf': "continuous_on (frontier (cball \<xi> r)) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
          and contg': "continuous_on (frontier (cball \<xi> r)) g"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
      by (blast intro: contf contg continuous_on_subset)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
    have froc: "frontier(cball \<xi> r) \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
      using \<open>0 < r\<close> by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
    moreover have "continuous_on (frontier (cball \<xi> r)) (norm o f)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
      using contf' continuous_on_compose continuous_on_norm_id by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
    ultimately obtain w where w: "w \<in> frontier(cball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
                          and now: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (f w) \<le> norm (f x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
      apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]])
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   332
      apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
    then have fw: "0 < norm (f w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
      by (simp add: fnz')
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
    have "continuous_on (frontier (cball \<xi> r)) (norm o g)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
      using contg' continuous_on_compose continuous_on_norm_id by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
    then obtain v where v: "v \<in> frontier(cball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
               and nov: "\<And>x. x \<in> frontier(cball \<xi> r) \<Longrightarrow> norm (g v) \<ge> norm (g x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
      apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]])
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   341
      apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
    then have fv: "0 < norm (f v)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
      by (simp add: fnz')
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
    have "norm ((deriv ^^ 0) g \<xi>) \<le> fact 0 * norm (g v) / r ^ 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
      by (rule Cauchy_inequality [OF holg contg \<open>0 < r\<close>]) (simp add: dist_norm nov)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
    then have "cmod (g \<xi>) \<le> norm (g v)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
      by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
    with w have wr: "norm (\<xi> - w) = r" and nfw: "norm (f w) \<le> norm (f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
      apply (simp_all add: dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
      by (metis \<open>0 < cmod (f \<xi>)\<close> g_def less_imp_inverse_less norm_inverse not_le now order_trans v)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
    with fw have False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
      using norm_less by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  }
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  with that show ?thesis by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
theorem open_mapping_thm:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
  assumes holf: "f holomorphic_on S"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   360
      and S: "open S" and "connected S"
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   361
      and "open U" and "U \<subseteq> S"
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
   362
      and fne: "\<not> f constant_on S"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
    shows "open (f ` U)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
  have *: "open (f ` U)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
          if "U \<noteq> {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\<And>x. \<exists>y \<in> U. f y \<noteq> x"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
          for U
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
  proof (clarsimp simp: open_contains_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
    fix \<xi> assume \<xi>: "\<xi> \<in> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
    show "\<exists>e>0. ball (f \<xi>) e \<subseteq> f ` U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
    proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
      have hol: "(\<lambda>z. f z - f \<xi>) holomorphic_on U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
        by (rule holomorphic_intros that)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
      obtain s where "0 < s" and sbU: "ball \<xi> s \<subseteq> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
                 and sne: "\<And>z. z \<in> ball \<xi> s - {\<xi>} \<Longrightarrow> (\<lambda>z. f z - f \<xi>) z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
        using isolated_zeros [OF hol U \<xi>]  by (metis fneU right_minus_eq)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
      obtain r where "0 < r" and r: "cball \<xi> r \<subseteq> ball \<xi> s"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
        apply (rule_tac r="s/2" in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
        using \<open>0 < s\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
      have "cball \<xi> r \<subseteq> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
        using sbU r by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
      then have frsbU: "frontier (cball \<xi> r) \<subseteq> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
        using Diff_subset frontier_def order_trans by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
      then have cof: "compact (frontier(cball \<xi> r))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
        by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
      have frne: "frontier (cball \<xi> r) \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
        using \<open>0 < r\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
      have contfr: "continuous_on (frontier (cball \<xi> r)) (\<lambda>z. norm (f z - f \<xi>))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
        apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
        using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
      obtain w where "norm (\<xi> - w) = r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
                 and w: "(\<And>z. norm (\<xi> - z) = r \<Longrightarrow> norm (f w - f \<xi>) \<le> norm(f z - f \<xi>))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
        apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
        apply (simp add: dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
        done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   396
      moreover define \<epsilon> where "\<epsilon> \<equiv> norm (f w - f \<xi>) / 3"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
      ultimately have "0 < \<epsilon>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
        using \<open>0 < r\<close> dist_complex_def r sne by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
      have "ball (f \<xi>) \<epsilon> \<subseteq> f ` U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
      proof
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
        fix \<gamma>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
        assume \<gamma>: "\<gamma> \<in> ball (f \<xi>) \<epsilon>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
        have *: "cmod (\<gamma> - f \<xi>) < cmod (\<gamma> - f z)" if "cmod (\<xi> - z) = r" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
        proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
          have lt: "cmod (f w - f \<xi>) / 3 < cmod (\<gamma> - f z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
            using w [OF that] \<gamma>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
            using dist_triangle2 [of "f \<xi>" "\<gamma>"  "f z"] dist_triangle2 [of "f \<xi>" "f z" \<gamma>]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
            by (simp add: \<epsilon>_def dist_norm norm_minus_commute)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
          show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
            by (metis \<epsilon>_def dist_commute dist_norm less_trans lt mem_ball \<gamma>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
       qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
       have "continuous_on (cball \<xi> r) (\<lambda>z. \<gamma> - f z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
          apply (rule continuous_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
          using \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
          apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
          done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
        moreover have "(\<lambda>z. \<gamma> - f z) holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
          apply (rule holomorphic_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
          apply (metis \<open>cball \<xi> r \<subseteq> U\<close> \<open>f holomorphic_on U\<close> holomorphic_on_subset interior_cball interior_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
          done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
        ultimately obtain z where "z \<in> ball \<xi> r" "\<gamma> - f z = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
          apply (rule holomorphic_contract_to_zero)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
          apply (blast intro!: \<open>0 < r\<close> *)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
          done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
        then show "\<gamma> \<in> f ` U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
          using \<open>cball \<xi> r \<subseteq> U\<close> by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
      qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
      then show ?thesis using  \<open>0 < \<epsilon>\<close> by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
  have "open (f ` X)" if "X \<in> components U" for X
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
    have holfU: "f holomorphic_on U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
      using \<open>U \<subseteq> S\<close> holf holomorphic_on_subset by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    have "X \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
      using that by (simp add: in_components_nonempty)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
    moreover have "open X"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
      using that \<open>open U\<close> open_components by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
    moreover have "connected X"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
      using that in_components_maximal by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
    moreover have "f holomorphic_on X"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
      by (meson that holfU holomorphic_on_subset in_components_maximal)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
    moreover have "\<exists>y\<in>X. f y \<noteq> x" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
    proof (rule ccontr)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
      assume not: "\<not> (\<exists>y\<in>X. f y \<noteq> x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
      have "X \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
        using \<open>U \<subseteq> S\<close> in_components_subset that by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
      obtain w where w: "w \<in> X" using \<open>X \<noteq> {}\<close> by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
      have wis: "w islimpt X"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
        using w \<open>open X\<close> interior_eq by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
      have hol: "(\<lambda>z. f z - x) holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
        by (simp add: holf holomorphic_on_diff)
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   453
      with fne [unfolded constant_on_def] 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   454
           analytic_continuation[OF hol S \<open>connected S\<close> \<open>X \<subseteq> S\<close> _ wis] not \<open>X \<subseteq> S\<close> w
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
      show False by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
    ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
      by (rule *)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  qed
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69712
diff changeset
   460
  then have "open (f ` \<Union>(components U))"
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62837
diff changeset
   461
    by (metis (no_types, lifting) imageE image_Union open_Union)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
  then show ?thesis
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62837
diff changeset
   463
    by force
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   466
text\<open>No need for \<^term>\<open>S\<close> to be connected. But the nonconstant condition is stronger.\<close>
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
   467
corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm2:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
      and S: "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
      and "open U" "U \<subseteq> S"
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
   471
      and fnc: "\<And>X. \<lbrakk>open X; X \<subseteq> S; X \<noteq> {}\<rbrakk> \<Longrightarrow> \<not> f constant_on X"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
    shows "open (f ` U)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
proof -
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62837
diff changeset
   474
  have "S = \<Union>(components S)" by simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
  with \<open>U \<subseteq> S\<close> have "U = (\<Union>C \<in> components S. C \<inter> U)" by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
  then have "f ` U = (\<Union>C \<in> components S. f ` (C \<inter> U))"
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62837
diff changeset
   477
    using image_UN by fastforce
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  moreover
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
  { fix C assume "C \<in> components S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
    with S \<open>C \<in> components S\<close> open_components in_components_connected
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
    have C: "open C" "connected C" by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
    have "C \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
      by (metis \<open>C \<in> components S\<close> in_components_maximal)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
    have nf: "\<not> f constant_on C"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
      apply (rule fnc)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
      using C \<open>C \<subseteq> S\<close> \<open>C \<in> components S\<close> in_components_nonempty by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
    have "f holomorphic_on C"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
      by (metis holf holomorphic_on_subset \<open>C \<subseteq> S\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
    then have "open (f ` (C \<inter> U))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
      apply (rule open_mapping_thm [OF _ C _ _ nf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
      apply (simp add: C \<open>open U\<close> open_Int, blast)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  } ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
    by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
   497
corollary\<^marker>\<open>tag unimportant\<close> open_mapping_thm3:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
      and "open S" and injf: "inj_on f S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
    shows  "open (f ` S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
apply (rule open_mapping_thm2 [OF holf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
using assms
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
apply (simp_all add:)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
using injective_not_constant subset_inj_on by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   506
subsection\<open>Maximum modulus principle\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   508
text\<open>If \<^term>\<open>f\<close> is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
   509
   properly within the domain of \<^term>\<open>f\<close>.\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
proposition maximum_modulus_principle:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
  assumes holf: "f holomorphic_on S"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   513
      and S: "open S" and "connected S"
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   514
      and "open U" and "U \<subseteq> S" and "\<xi> \<in> U"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
      and no: "\<And>z. z \<in> U \<Longrightarrow> norm(f z) \<le> norm(f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
    shows "f constant_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
proof (rule ccontr)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
  assume "\<not> f constant_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
  then have "open (f ` U)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
    using open_mapping_thm assms by blast
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
   521
  moreover have "\<not> open (f ` U)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
    have "\<exists>t. cmod (f \<xi> - t) < e \<and> t \<notin> f ` U" if "0 < e" for e
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
      apply (rule_tac x="if 0 < Re(f \<xi>) then f \<xi> + (e/2) else f \<xi> - (e/2)" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
      using that
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
      apply (simp add: dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
      apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
      unfolding open_contains_ball by (metis \<open>\<xi> \<in> U\<close> contra_subsetD dist_norm imageI mem_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
  ultimately show False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
    by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
proposition maximum_modulus_frontier:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
  assumes holf: "f holomorphic_on (interior S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
      and contf: "continuous_on (closure S) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
      and bos: "bounded S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
      and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> norm(f z) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
      and "\<xi> \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    shows "norm(f \<xi>) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
  have "compact (closure S)" using bos
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
    by (simp add: bounded_closure compact_eq_bounded_closed)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
  moreover have "continuous_on (closure S) (cmod \<circ> f)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
    using contf continuous_on_compose continuous_on_norm_id by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
  ultimately obtain z where zin: "z \<in> closure S" and z: "\<And>y. y \<in> closure S \<Longrightarrow> (cmod \<circ> f) y \<le> (cmod \<circ> f) z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
    using continuous_attains_sup [of "closure S" "norm o f"] \<open>\<xi> \<in> S\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
  then consider "z \<in> frontier S" | "z \<in> interior S" using frontier_def by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
  then have "norm(f z) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
  proof cases
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
    case 1 then show ?thesis using leB by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
    case 2
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
    have zin: "z \<in> connected_component_set (interior S) z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
      by (simp add: 2)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
    have "f constant_on (connected_component_set (interior S) z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
      apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
      apply (metis connected_component_subset holf holomorphic_on_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
      apply (simp_all add: open_connected_component)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
      by (metis closure_subset comp_eq_dest_lhs  interior_subset subsetCE z connected_component_in)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
    then obtain c where c: "\<And>w. w \<in> connected_component_set (interior S) z \<Longrightarrow> f w = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
      by (auto simp: constant_on_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
    have "f ` closure(connected_component_set (interior S) z) \<subseteq> {c}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
      apply (rule image_closure_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
      apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
      using c
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
      apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
    then have cc: "\<And>w. w \<in> closure(connected_component_set (interior S) z) \<Longrightarrow> f w = c" by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
    have "frontier(connected_component_set (interior S) z) \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
      apply (simp add: frontier_eq_empty)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
      by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
    then obtain w where w: "w \<in> frontier(connected_component_set (interior S) z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
       by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
    then have "norm (f z) = norm (f w)"  by (simp add: "2" c cc frontier_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
    also have "... \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
      apply (rule leB)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
      using w
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
using frontier_interior_subset frontier_of_connected_component_subset by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
    finally show ?thesis .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
    using z \<open>\<xi> \<in> S\<close> closure_subset by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
   588
corollary\<^marker>\<open>tag unimportant\<close> maximum_real_frontier:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
  assumes holf: "f holomorphic_on (interior S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
      and contf: "continuous_on (closure S) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
      and bos: "bounded S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
      and leB: "\<And>z. z \<in> frontier S \<Longrightarrow> Re(f z) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
      and "\<xi> \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
    shows "Re(f \<xi>) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
using maximum_modulus_frontier [of "exp o f" S "exp B"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
      Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
   599
subsection\<^marker>\<open>tag unimportant\<close> \<open>Factoring out a zero according to its order\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
lemma holomorphic_factor_order_of_zero:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
      and os: "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
      and "\<xi> \<in> S" "0 < n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
      and dnz: "(deriv ^^ n) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
      and dfz: "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
   obtains g r where "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
                "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
                "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
                "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
  obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
  then have holfb: "f holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
    using holf holomorphic_on_subset by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   615
  define g where "g w = suminf (\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i)" for w
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
  have sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
   and feq: "f w - f \<xi> = (w - \<xi>)^n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
       if w: "w \<in> ball \<xi> r" for w
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
  proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   620
    define powf where "powf = (\<lambda>i. (deriv ^^ i) f \<xi>/(fact i) * (w - \<xi>)^i)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
    have sing: "{..<n} - {i. powf i = 0} = (if f \<xi> = 0 then {} else {0})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
      unfolding powf_def using \<open>0 < n\<close> dfz by (auto simp: dfz; metis funpow_0 not_gr0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
    have "powf sums f w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
      unfolding powf_def by (rule holomorphic_power_series [OF holfb w])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
    moreover have "(\<Sum>i<n. powf i) = f \<xi>"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
   626
      apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric])
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   627
      apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
      apply (simp only: dfz sing)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
      apply (simp add: powf_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
    ultimately have fsums: "(\<lambda>i. powf (i+n)) sums (f w - f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
      using w sums_iff_shift' by metis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
    then have *: "summable (\<lambda>i. (w - \<xi>) ^ n * ((deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n)))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
      unfolding powf_def using sums_summable
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
      by (auto simp: power_add mult_ac)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
    have "summable (\<lambda>i. (deriv ^^ (i + n)) f \<xi> * (w - \<xi>) ^ i / fact (i + n))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
    proof (cases "w=\<xi>")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
      case False then show ?thesis
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   639
        using summable_mult [OF *, of "1 / (w - \<xi>) ^ n"] by simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
    next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
      case True then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
        by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
                 split: if_split_asm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
    then show sumsg: "(\<lambda>i. (deriv ^^ (i + n)) f \<xi> / (fact(i + n)) * (w - \<xi>)^i) sums g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
      by (simp add: summable_sums_iff g_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
    show "f w - f \<xi> = (w - \<xi>)^n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
      apply (rule sums_unique2)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
      apply (rule fsums [unfolded powf_def])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
      using sums_mult [OF sumsg, of "(w - \<xi>) ^ n"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
      by (auto simp: power_add mult_ac)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
  then have holg: "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
    by (meson sumsg power_series_holomorphic)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
  then have contg: "continuous_on (ball \<xi> r) g"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
    by (blast intro: holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
  have "g \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
    using dnz unfolding g_def
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
    by (subst suminf_finite [of "{0}"]) auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
  obtain d where "0 < d" and d: "\<And>w. w \<in> ball \<xi> d \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
    apply (rule exE [OF continuous_on_avoid [OF contg _ \<open>g \<xi> \<noteq> 0\<close>]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
    using \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
    apply force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
    by (metis \<open>0 < r\<close> less_trans mem_ball not_less_iff_gr_or_eq)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
    apply (rule that [where g=g and r ="min r d"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
    using \<open>0 < r\<close> \<open>0 < d\<close> holg
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
    apply (auto simp: feq holomorphic_on_subset subset_ball d)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
lemma holomorphic_factor_order_of_zero_strong:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
  assumes holf: "f holomorphic_on S" "open S"  "\<xi> \<in> S" "0 < n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
      and "(deriv ^^ n) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
      and "\<And>i. \<lbrakk>0 < i; i < n\<rbrakk> \<Longrightarrow> (deriv ^^ i) f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
   obtains g r where "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
                "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
                "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
                "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
  obtain g r where "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
               and holg: "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
               and feq: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>)^n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
               and gne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
    by (auto intro: holomorphic_factor_order_of_zero [OF assms])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
  have con: "continuous_on (ball \<xi> r) (\<lambda>z. deriv g z / g z)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
    by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   689
  have cd: "\<And>x. dist \<xi> x < r \<Longrightarrow> (\<lambda>z. deriv g z / g z) field_differentiable at x"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
    apply (rule derivative_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
    using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
   692
    apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
    using gne mem_ball by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
  obtain h where h: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> (h has_field_derivative deriv g x / g x) (at x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
    apply (rule exE [OF holomorphic_convex_primitive [of "ball \<xi> r" "{}" "\<lambda>z. deriv g z / g z"]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
    apply (auto simp: con cd)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
    apply (metis open_ball at_within_open mem_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
  then have "continuous_on (ball \<xi> r) h"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
    by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
  then have con: "continuous_on (ball \<xi> r) (\<lambda>x. exp (h x) / g x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
    by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
  have 0: "dist \<xi> x < r \<Longrightarrow> ((\<lambda>x. exp (h x) / g x) has_field_derivative 0) (at x)" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
    apply (rule h derivative_eq_intros | simp)+
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   705
    apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
    using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
  obtain c where c: "\<And>x. x \<in> ball \<xi> r \<Longrightarrow> exp (h x) / g x = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
    by (rule DERIV_zero_connected_constant [of "ball \<xi> r" "{}" "\<lambda>x. exp(h x) / g x"]) (auto simp: con 0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
  have hol: "(\<lambda>z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
    apply (rule holomorphic_on_compose [unfolded o_def, where g = exp])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
    apply (rule holomorphic_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
    using h holomorphic_on_open apply blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
    apply (rule holomorphic_intros)+
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   715
    using \<open>0 < n\<close> apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
    apply (rule holomorphic_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
    apply (rule that [where g="\<lambda>z. exp((Ln(inverse c) + h z)/n)" and r =r])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
    using \<open>0 < r\<close> \<open>0 < n\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
    apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
    apply (rule hol)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
    apply (simp add: Transcendental.exp_add gne)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
lemma
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
  fixes k :: "'a::wellorder"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
  assumes a_def: "a == LEAST x. P x" and P: "P k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
  shows def_LeastI: "P a" and def_Least_le: "a \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
unfolding a_def
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
by (rule LeastI Least_le; rule P)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
lemma holomorphic_factor_zero_nonconstant:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
  assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
      and "\<xi> \<in> S" "f \<xi> = 0"
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
   738
      and nonconst: "\<not> f constant_on S"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
   obtains g r n
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
      where "0 < n"  "0 < r"  "ball \<xi> r \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
            "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
            "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w = (w - \<xi>)^n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
            "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
  case True then show ?thesis
66660
bc3584f7ac0c Using the "constant_on" operator
paulson <lp15@cam.ac.uk>
parents: 66486
diff changeset
   746
    using holomorphic_fun_eq_const_on_connected [OF holf S _ \<open>\<xi> \<in> S\<close>] nonconst by (simp add: constant_on_def)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
  case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
  then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0" by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
  obtain r0 where "r0 > 0" "ball \<xi> r0 \<subseteq> S" using S openE \<open>\<xi> \<in> S\<close> by auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   751
  define n where "n \<equiv> LEAST n. (deriv ^^ n) f \<xi> \<noteq> 0"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
  have n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
    by (rule def_LeastI [OF n_def]) (rule n0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
  then have "0 < n" using \<open>f \<xi> = 0\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
    using funpow_0 by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
  have n_min: "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
    using def_Least_le [OF n_def] not_le by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
  then obtain g r1
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
    where  "0 < r1" "g holomorphic_on ball \<xi> r1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
           "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> f w = (w - \<xi>) ^ n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
           "\<And>w. w \<in> ball \<xi> r1 \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
    by (auto intro: holomorphic_factor_order_of_zero [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne] simp: \<open>f \<xi> = 0\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
    apply (rule_tac g=g and r="min r0 r1" and n=n in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
    using \<open>0 < n\<close> \<open>0 < r0\<close> \<open>0 < r1\<close> \<open>ball \<xi> r0 \<subseteq> S\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
    apply (auto simp: subset_ball intro: holomorphic_on_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
lemma holomorphic_lower_bound_difference:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
  assumes holf: "f holomorphic_on S" and S: "open S" "connected S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
      and "\<xi> \<in> S" and "\<phi> \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
      and fne: "f \<phi> \<noteq> f \<xi>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
   obtains k n r
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
      where "0 < k"  "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
            "ball \<xi> r \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
            "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> k * norm(w - \<xi>)^n \<le> norm(f w - f \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   780
  define n where "n = (LEAST n. 0 < n \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
  obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
    using fne holomorphic_fun_eq_const_on_connected [OF holf S] \<open>\<xi> \<in> S\<close> \<open>\<phi> \<in> S\<close> by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
  then have "0 < n" and n_ne: "(deriv ^^ n) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
    unfolding n_def by (metis (mono_tags, lifting) LeastI)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
  have n_min: "\<And>k. \<lbrakk>0 < k; k < n\<rbrakk> \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
    unfolding n_def by (blast dest: not_less_Least)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  then obtain g r
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
    where "0 < r" and holg: "g holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
      and fne: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> f w - f \<xi> = (w - \<xi>) ^ n * g w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
      and gnz: "\<And>w. w \<in> ball \<xi> r \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
      by (auto intro: holomorphic_factor_order_of_zero  [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> \<open>n > 0\<close> n_ne])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
  obtain e where "e>0" and e: "ball \<xi> e \<subseteq> S" using assms by (blast elim!: openE)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
  then have holfb: "f holomorphic_on ball \<xi> e"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
    using holf holomorphic_on_subset by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   795
  define d where "d = (min e r) / 2"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
  have "0 < d" using \<open>0 < r\<close> \<open>0 < e\<close> by (simp add: d_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
  have "d < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
    using \<open>0 < r\<close> by (auto simp: d_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
  then have cbb: "cball \<xi> d \<subseteq> ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
    by (auto simp: cball_subset_ball_iff)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
  then have "g holomorphic_on cball \<xi> d"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
    by (rule holomorphic_on_subset [OF holg])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
  then have "closed (g ` cball \<xi> d)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
    by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
  moreover have "g ` cball \<xi> d \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
    using \<open>0 < d\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
  ultimately obtain x where x: "x \<in> g ` cball \<xi> d" and "\<And>y. y \<in> g ` cball \<xi> d \<Longrightarrow> dist 0 x \<le> dist 0 y"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
    by (rule distance_attains_inf) blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
  then have leg: "\<And>w. w \<in> cball \<xi> d \<Longrightarrow> norm x \<le> norm (g w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
    by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
  have "ball \<xi> d \<subseteq> cball \<xi> d" by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
  also have "... \<subseteq> ball \<xi> e" using \<open>0 < d\<close> d_def by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
  also have "... \<subseteq> S" by (rule e)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
  finally have dS: "ball \<xi> d \<subseteq> S" .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
  moreover have "x \<noteq> 0" using gnz x \<open>d < r\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
  ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
    apply (rule_tac k="norm x" and n=n and r=d in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
    using \<open>d < r\<close> leg
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
    apply (auto simp: \<open>0 < d\<close> fne norm_mult norm_power algebra_simps mult_right_mono)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
lemma
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
  assumes holf: "f holomorphic_on (S - {\<xi>})" and \<xi>: "\<xi> \<in> interior S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
    shows holomorphic_on_extend_lim:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
          "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
           ((\<lambda>z. (z - \<xi>) * f z) \<longlongrightarrow> 0) (at \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
          (is "?P = ?Q")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
     and holomorphic_on_extend_bounded:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
          "(\<exists>g. g holomorphic_on S \<and> (\<forall>z \<in> S - {\<xi>}. g z = f z)) \<longleftrightarrow>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
           (\<exists>B. eventually (\<lambda>z. norm(f z) \<le> B) (at \<xi>))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
          (is "?P = ?R")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
  obtain \<delta> where "0 < \<delta>" and \<delta>: "ball \<xi> \<delta> \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
    using \<xi> mem_interior by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
  have "?R" if holg: "g holomorphic_on S" and gf: "\<And>z. z \<in> S - {\<xi>} \<Longrightarrow> g z = f z" for g
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
    have *: "\<forall>\<^sub>F z in at \<xi>. dist (g z) (g \<xi>) < 1 \<longrightarrow> cmod (f z) \<le> cmod (g \<xi>) + 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
      apply (simp add: eventually_at)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
      apply (rule_tac x="\<delta>" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
      using \<delta> \<open>0 < \<delta>\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
      apply (clarsimp simp:)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
      apply (drule_tac c=x in subsetD)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
      apply (simp add: dist_commute)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
      by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
    have "continuous_on (interior S) g"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
      by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
    then have "\<And>x. x \<in> interior S \<Longrightarrow> (g \<longlongrightarrow> g x) (at x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
      using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
    then have "(g \<longlongrightarrow> g \<xi>) (at \<xi>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
      by (simp add: \<xi>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
      apply (rule_tac x="norm(g \<xi>) + 1" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
      apply (rule eventually_mp [OF * tendstoD [where e=1]], auto)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
  moreover have "?Q" if "\<forall>\<^sub>F z in at \<xi>. cmod (f z) \<le> B" for B
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
    by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
  moreover have "?P" if "(\<lambda>z. (z - \<xi>) * f z) \<midarrow>\<xi>\<rightarrow> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
  proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   861
    define h where [abs_def]: "h z = (z - \<xi>)^2 * f z" for z
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
    have h0: "(h has_field_derivative 0) (at \<xi>)"
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   863
      apply (simp add: h_def has_field_derivative_iff)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
      apply (rule Lim_transform_within [OF that, of 1])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
      apply (auto simp: divide_simps power2_eq_square)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
    have holh: "h holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
    proof (simp add: holomorphic_on_def, clarify)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
      fix z assume "z \<in> S"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   870
      show "h field_differentiable at z within S"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
      proof (cases "z = \<xi>")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
        case True then show ?thesis
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   873
          using field_differentiable_at_within field_differentiable_def h0 by blast
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
      next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
        case False
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   876
        then have "f field_differentiable at z within S"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
          using holomorphic_onD [OF holf, of z] \<open>z \<in> S\<close>
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
   878
          unfolding field_differentiable_def has_field_derivative_iff
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
          by (force intro: exI [where x="dist \<xi> z"] elim: Lim_transform_within_set [unfolded eventually_at])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
        then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
          by (simp add: h_def power2_eq_square derivative_intros)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
      qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
    qed
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
   884
    define g where [abs_def]: "g z = (if z = \<xi> then deriv h \<xi> else (h z - h \<xi>) / (z - \<xi>))" for z
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
    have holg: "g holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
      unfolding g_def by (rule pole_lemma [OF holh \<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
    show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
      apply (rule_tac x="\<lambda>z. if z = \<xi> then deriv g \<xi> else (g z - g \<xi>)/(z - \<xi>)" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
      apply (rule conjI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
      apply (rule pole_lemma [OF holg \<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
      apply (auto simp: g_def power2_eq_square divide_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
      using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
  ultimately show "?P = ?Q" and "?P = ?R"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
    by meson+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
   899
lemma pole_at_infinity:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
  assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \<longlongrightarrow> l) at_infinity"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
  obtains a n where "\<And>z. f z = (\<Sum>i\<le>n. a i * z^i)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
proof (cases "l = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
  case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  with tendsto_inverse [OF lim] show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
    apply (rule_tac a="(\<lambda>n. inverse l)" and n=0 in that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
    apply (simp add: Liouville_weak [OF holf, of "inverse l"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
  case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
  then have [simp]: "l = 0" .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
  proof (cases "\<exists>r. 0 < r \<and> (\<forall>z \<in> ball 0 r - {0}. f(inverse z) \<noteq> 0)")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
    case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
      then obtain r where "0 < r" and r: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> f(inverse z) \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
             by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
      have 1: "inverse \<circ> f \<circ> inverse holomorphic_on ball 0 r - {0}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
        by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
      have 2: "0 \<in> interior (ball 0 r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
        using \<open>0 < r\<close> by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
      have "\<exists>B. 0<B \<and> eventually (\<lambda>z. cmod ((inverse \<circ> f \<circ> inverse) z) \<le> B) (at 0)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
        apply (rule exI [where x=1])
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   922
        apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
        using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
        apply (rule eventually_mono)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
        apply (simp add: dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
        done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
      with holomorphic_on_extend_bounded [OF 1 2]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
      obtain g where holg: "g holomorphic_on ball 0 r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
                 and geq: "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> g z = (inverse \<circ> f \<circ> inverse) z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
        by meson
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
      have ifi0: "(inverse \<circ> f \<circ> inverse) \<midarrow>0\<rightarrow> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
        using \<open>l = 0\<close> lim lim_at_infinity_0 by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
      have g2g0: "g \<midarrow>0\<rightarrow> g 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
        using \<open>0 < r\<close> centre_in_ball continuous_at continuous_on_eq_continuous_at holg
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
        by (blast intro: holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
      have g2g1: "g \<midarrow>0\<rightarrow> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
        apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
        using \<open>0 < r\<close> by (auto simp: geq)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
      have [simp]: "g 0 = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
        by (rule tendsto_unique [OF _ g2g0 g2g1]) simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
      have "ball 0 r - {0::complex} \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
        using \<open>0 < r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
        apply (clarsimp simp: ball_def dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
        apply (drule_tac c="of_real r/2" in subsetD, auto)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
        done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
      then obtain w::complex where "w \<noteq> 0" and w: "norm w < r" by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
      then have "g w \<noteq> 0" by (simp add: geq r)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
      obtain B n e where "0 < B" "0 < e" "e \<le> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
                     and leg: "\<And>w. norm w < e \<Longrightarrow> B * cmod w ^ n \<le> cmod (g w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
        apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
        using \<open>0 < r\<close> w \<open>g w \<noteq> 0\<close> by (auto simp: ball_subset_ball_iff)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
      have "cmod (f z) \<le> cmod z ^ n / B" if "2/e \<le> cmod z" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
      proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
        have ize: "inverse z \<in> ball 0 e - {0}" using that \<open>0 < e\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
          by (auto simp: norm_divide divide_simps algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
        then have [simp]: "z \<noteq> 0" and izr: "inverse z \<in> ball 0 r - {0}" using  \<open>e \<le> r\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
          by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
        then have [simp]: "f z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
          using r [of "inverse z"] by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
        have [simp]: "f z = inverse (g (inverse z))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
          using izr geq [of "inverse z"] by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
        show ?thesis using ize leg [of "inverse z"]  \<open>0 < B\<close>  \<open>0 < e\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
          by (simp add: divide_simps norm_divide algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
      qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
      then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
        apply (rule_tac a = "\<lambda>k. (deriv ^^ k) f 0 / (fact k)" and n=n in that)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
   967
        apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
        done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
    then have fi0: "\<And>r. r > 0 \<Longrightarrow> \<exists>z\<in>ball 0 r - {0}. f (inverse z) = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
      by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
    have fz0: "f z = 0" if "0 < r" and lt1: "\<And>x. x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> inverse (cmod (f (inverse x))) < 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
              for z r
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
    proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
      have f0: "(f \<longlongrightarrow> 0) at_infinity"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
      proof -
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67399
diff changeset
   978
        have DIM_complex[intro]: "2 \<le> DIM(complex)"  \<comment> \<open>should not be necessary!\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
          by simp
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   980
        from lt1 have "f (inverse x) \<noteq> 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> 1 < cmod (f (inverse x))" for x
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   981
          using one_less_inverse by force
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   982
        then have **: "cmod (f (inverse x)) \<le> 1 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> cmod x < r \<Longrightarrow> f (inverse x) = 0" for x
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   983
          by force
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   984
        then have *: "(f \<circ> inverse) ` (ball 0 r - {0}) \<subseteq> {0} \<union> - ball 0 1"
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   985
          by force
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
        have "continuous_on (inverse ` (ball 0 r - {0})) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
          using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
        then have "connected ((f \<circ> inverse) ` (ball 0 r - {0}))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
          apply (intro connected_continuous_image continuous_intros)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
          apply (force intro: connected_punctured_ball)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
          done
69661
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   992
        then have "{0} \<inter> (f \<circ> inverse) ` (ball 0 r - {0}) = {} \<or> - ball 0 1 \<inter> (f \<circ> inverse) ` (ball 0 r - {0}) = {}"
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   993
          by (rule connected_closedD) (use * in auto)
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   994
        then have "w \<noteq> 0 \<Longrightarrow> cmod w < r \<Longrightarrow> f (inverse w) = 0" for w
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   995
          using fi0 **[of w] \<open>0 < r\<close>
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   996
          apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le)
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   997
           apply fastforce
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   998
          apply (drule bspec [of _ _ w])
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
   999
           apply (auto dest: less_imp_le)
a03a63b81f44 tuned proofs
haftmann
parents: 69597
diff changeset
  1000
          done
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
        then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
          apply (simp add: lim_at_infinity_0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
          apply (rule Lim_eventually)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
          apply (simp add: eventually_at)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
          apply (rule_tac x=r in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
          apply (simp add: \<open>0 < r\<close> dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
          done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
      qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
      obtain w where "w \<in> ball 0 r - {0}" and "f (inverse w) = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
        using False \<open>0 < r\<close> by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
      then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
        by (auto simp: f0 Liouville_weak [OF holf, of 0])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
    show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
      apply (rule that [of "\<lambda>n. 0" 0])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
      using lim [unfolded lim_at_infinity_0]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
      apply (simp add: Lim_at dist_norm norm_inverse)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
      apply (drule_tac x=1 in spec)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
      using fz0 apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
  1024
subsection\<^marker>\<open>tag unimportant\<close> \<open>Entire proper functions are precisely the non-trivial polynomials\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1026
lemma proper_map_polyfun:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
    fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
  assumes "closed S" and "compact K" and c: "c i \<noteq> 0" "1 \<le> i" "i \<le> n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
    shows "compact (S \<inter> {z. (\<Sum>i\<le>n. c i * z^i) \<in> K})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
  obtain B where "B > 0" and B: "\<And>x. x \<in> K \<Longrightarrow> norm x \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
    by (metis compact_imp_bounded \<open>compact K\<close> bounded_pos)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
  have *: "norm x \<le> b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
            if "\<And>x. b \<le> norm x \<Longrightarrow> B + 1 \<le> norm (\<Sum>i\<le>n. c i * x ^ i)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
               "(\<Sum>i\<le>n. c i * x ^ i) \<in> K"  for b x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
    have "norm (\<Sum>i\<le>n. c i * x ^ i) \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
      using B that by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
    moreover have "\<not> B + 1 \<le> B"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
      by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
    ultimately show "norm x \<le> b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
      using that by (metis (no_types) less_eq_real_def not_less order_trans)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
  have "bounded {z. (\<Sum>i\<le>n. c i * z ^ i) \<in> K}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
    using polyfun_extremal [where c=c and B="B+1", OF c]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
    by (auto simp: bounded_pos eventually_at_infinity_pos *)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1047
  moreover have "closed ((\<lambda>z. (\<Sum>i\<le>n. c i * z ^ i)) -` K)"
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1048
    apply (intro allI continuous_closed_vimage continuous_intros)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
    using \<open>compact K\<close> compact_eq_bounded_closed by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
  ultimately show ?thesis
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1051
    using closed_Int_compact [OF \<open>closed S\<close>] compact_eq_bounded_closed
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1052
    by (auto simp add: vimage_def)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1055
lemma proper_map_polyfun_univ:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
    fixes c :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,heine_borel}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
  assumes "compact K" "c i \<noteq> 0" "1 \<le> i" "i \<le> n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
    shows "compact ({z. (\<Sum>i\<le>n. c i * z^i) \<in> K})"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1059
  using proper_map_polyfun [of UNIV K c i n] assms by simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1061
lemma proper_map_polyfun_eq:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
  assumes "f holomorphic_on UNIV"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
    shows "(\<forall>k. compact k \<longrightarrow> compact {z. f z \<in> k}) \<longleftrightarrow>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
           (\<exists>c n. 0 < n \<and> (c n \<noteq> 0) \<and> f = (\<lambda>z. \<Sum>i\<le>n. c i * z^i))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
          (is "?lhs = ?rhs")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
proof
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
  assume compf [rule_format]: ?lhs
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
  have 2: "\<exists>k. 0 < k \<and> a k \<noteq> 0 \<and> f = (\<lambda>z. \<Sum>i \<le> k. a i * z ^ i)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
        if "\<And>z. f z = (\<Sum>i\<le>n. a i * z ^ i)" for a n
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
  proof (cases "\<forall>i\<le>n. 0<i \<longrightarrow> a i = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
    case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
    then have [simp]: "\<And>z. f z = a 0"
70113
c8deb8ba6d05 Fixing the main Homology theory; also moving a lot of sum/prod lemmas into their generic context
paulson <lp15@cam.ac.uk>
parents: 70065
diff changeset
  1073
      by (simp add: that sum.atMost_shift)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
    have False using compf [of "{a 0}"] by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
    then show ?thesis ..
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
    then obtain k where k: "0 < k" "k\<le>n" "a k \<noteq> 0" by force
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1079
    define m where "m = (GREATEST k. k\<le>n \<and> a k \<noteq> 0)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
    have m: "m\<le>n \<and> a m \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
      unfolding m_def
65965
088c79b40156 tuned names
nipkow
parents: 65963
diff changeset
  1082
      apply (rule GreatestI_nat [where b = n])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
      using k apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
    have [simp]: "a i = 0" if "m < i" "i \<le> n" for i
65965
088c79b40156 tuned names
nipkow
parents: 65963
diff changeset
  1086
      using Greatest_le_nat [where b = "n" and P = "\<lambda>k. k\<le>n \<and> a k \<noteq> 0"]
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
      using m_def not_le that by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
    have "k \<le> m"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
      unfolding m_def
65965
088c79b40156 tuned names
nipkow
parents: 65963
diff changeset
  1090
      apply (rule Greatest_le_nat [where b = "n"])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
      using k apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
    with k m show ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  1094
      by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
  have "((inverse \<circ> f) \<longlongrightarrow> 0) at_infinity"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
  proof (rule Lim_at_infinityI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
    fix e::real assume "0 < e"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
    with compf [of "cball 0 (inverse e)"]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
    show "\<exists>B. \<forall>x. B \<le> cmod x \<longrightarrow> dist ((inverse \<circ> f) x) 0 \<le> e"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1101
      apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
      apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
      apply (rule_tac x="b+1" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
      apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
  then show ?rhs
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
    apply (rule pole_at_infinity [OF assms])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
    using 2 apply blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
  assume ?rhs
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
  then obtain c n where "0 < n" "c n \<noteq> 0" "f = (\<lambda>z. \<Sum>i\<le>n. c i * z ^ i)" by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
  then have "compact {z. f z \<in> k}" if "compact k" for k
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
    by (auto intro: proper_map_polyfun_univ [OF that])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
  then show ?lhs by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1119
subsection \<open>Relating invertibility and nonvanishing of derivative\<close>
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1121
lemma has_complex_derivative_locally_injective:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
      and S: "\<xi> \<in> S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
      and dnz: "deriv f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
  obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
  have *: "\<exists>d>0. \<forall>x. dist \<xi> x < d \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) < e" if "e > 0" for e
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
    have contdf: "continuous_on S (deriv f)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
      by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \<open>open S\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
    obtain \<delta> where "\<delta>>0" and \<delta>: "\<And>x. \<lbrakk>x \<in> S; dist x \<xi> \<le> \<delta>\<rbrakk> \<Longrightarrow> cmod (deriv f x - deriv f \<xi>) \<le> e/2"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
      using continuous_onE [OF contdf \<open>\<xi> \<in> S\<close>, of "e/2"] \<open>0 < e\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
      by (metis dist_complex_def half_gt_zero less_imp_le)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
    obtain \<epsilon> where "\<epsilon>>0" "ball \<xi> \<epsilon> \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
      by (metis openE [OF \<open>open S\<close> \<open>\<xi> \<in> S\<close>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
    with \<open>\<delta>>0\<close> have "\<exists>\<delta>>0. \<forall>x. dist \<xi> x < \<delta> \<longrightarrow> onorm (\<lambda>v. deriv f x * v - deriv f \<xi> * v) \<le> e/2"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
      apply (rule_tac x="min \<delta> \<epsilon>" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
      apply (intro conjI allI impI Operator_Norm.onorm_le)
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66827
diff changeset
  1139
      apply simp
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
      apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
      apply (rule mult_right_mono [OF \<delta>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
      apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \<delta>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
    with \<open>e>0\<close> show ?thesis by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
  qed
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68629
diff changeset
  1146
  have "inj ((*) (deriv f \<xi>))"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
    using dnz by simp
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68629
diff changeset
  1148
  then obtain g' where g': "linear g'" "g' \<circ> (*) (deriv f \<xi>) = id"
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68629
diff changeset
  1149
    using linear_injective_left_inverse [of "(*) (deriv f \<xi>)"]
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
    by (auto simp: linear_times)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
    apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\<lambda>z h. deriv f z * h" and g' = g'])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
    using g' *
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
    apply (simp_all add: linear_conv_bounded_linear that)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1155
    using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
        holomorphic_on_imp_differentiable_at \<open>open S\<close> apply blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1160
lemma has_complex_derivative_locally_invertible:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
      and S: "\<xi> \<in> S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
      and dnz: "deriv f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
  obtains r where "r > 0" "ball \<xi> r \<subseteq> S" "open (f `  (ball \<xi> r))" "inj_on f (ball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
  obtain r where "r > 0" "ball \<xi> r \<subseteq> S" "inj_on f (ball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
    by (blast intro: that has_complex_derivative_locally_injective [OF assms])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
  then have \<xi>: "\<xi> \<in> ball \<xi> r" by simp
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
  1169
  then have nc: "\<not> f constant_on ball \<xi> r"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
    using \<open>inj_on f (ball \<xi> r)\<close> injective_not_constant by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
  have holf': "f holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
    using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
  have "open (f ` ball \<xi> r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
    apply (rule open_mapping_thm [OF holf'])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
    using nc apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
    using \<open>0 < r\<close> \<open>ball \<xi> r \<subseteq> S\<close> \<open>inj_on f (ball \<xi> r)\<close> that  by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1181
lemma holomorphic_injective_imp_regular:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
      and "open S" and injf: "inj_on f S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
      and "\<xi> \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
    shows "deriv f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
  obtain r where "r>0" and r: "ball \<xi> r \<subseteq> S" using assms by (blast elim!: openE)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
  have holf': "f holomorphic_on ball \<xi> r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
    using \<open>ball \<xi> r \<subseteq> S\<close> holf holomorphic_on_subset by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
  proof (cases "\<forall>n>0. (deriv ^^ n) f \<xi> = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
    case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
    have fcon: "f w = f \<xi>" if "w \<in> ball \<xi> r" for w
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
      apply (rule holomorphic_fun_eq_const_on_connected [OF holf'])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
      using True \<open>0 < r\<close> that by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
    have False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
      using fcon [of "\<xi> + r/2"] \<open>0 < r\<close> r injf unfolding inj_on_def
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
      by (metis \<open>\<xi> \<in> S\<close> contra_subsetD dist_commute fcon mem_ball perfect_choose_dist)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
    then show ?thesis ..
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
    then obtain n0 where n0: "n0 > 0 \<and> (deriv ^^ n0) f \<xi> \<noteq> 0" by blast
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1203
    define n where [abs_def]: "n = (LEAST n. n > 0 \<and> (deriv ^^ n) f \<xi> \<noteq> 0)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
    have n_ne: "n > 0" "(deriv ^^ n) f \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
      using def_LeastI [OF n_def n0] by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
    have n_min: "\<And>k. 0 < k \<Longrightarrow> k < n \<Longrightarrow> (deriv ^^ k) f \<xi> = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
      using def_Least_le [OF n_def] not_le by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
    obtain g \<delta> where "0 < \<delta>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
             and holg: "g holomorphic_on ball \<xi> \<delta>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
             and fd: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> f w - f \<xi> = ((w - \<xi>) * g w) ^ n"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
             and gnz: "\<And>w. w \<in> ball \<xi> \<delta> \<Longrightarrow> g w \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
      apply (rule holomorphic_factor_order_of_zero_strong [OF holf \<open>open S\<close> \<open>\<xi> \<in> S\<close> n_ne])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
      apply (blast intro: n_min)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
    show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
    proof (cases "n=1")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
      case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
      with n_ne show ?thesis by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
    next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
      case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
      have holgw: "(\<lambda>w. (w - \<xi>) * g w) holomorphic_on ball \<xi> (min r \<delta>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
        apply (rule holomorphic_intros)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
        using holg by (simp add: holomorphic_on_subset subset_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
      have gd: "\<And>w. dist \<xi> w < \<delta> \<Longrightarrow> (g has_field_derivative deriv g w) (at w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
        using holg
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1226
        by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
      have *: "\<And>w. w \<in> ball \<xi> (min r \<delta>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
            \<Longrightarrow> ((\<lambda>w. (w - \<xi>) * g w) has_field_derivative ((w - \<xi>) * deriv g w + g w))
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
                (at w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
        by (rule gd derivative_eq_intros | simp)+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
      have [simp]: "deriv (\<lambda>w. (w - \<xi>) * g w) \<xi> \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
        using * [of \<xi>] \<open>0 < \<delta>\<close> \<open>0 < r\<close> by (simp add: DERIV_imp_deriv gnz)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
      obtain T where "\<xi> \<in> T" "open T" and Tsb: "T \<subseteq> ball \<xi> (min r \<delta>)" and oimT: "open ((\<lambda>w. (w - \<xi>) * g w) ` T)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
        apply (rule has_complex_derivative_locally_invertible [OF holgw, of \<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
        using \<open>0 < r\<close> \<open>0 < \<delta>\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
        apply (simp_all add:)
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  1237
        by (meson open_ball centre_in_ball)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1238
      define U where "U = (\<lambda>w. (w - \<xi>) * g w) ` T"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
      have "open U" by (metis oimT U_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
      have "0 \<in> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
        apply (auto simp: U_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
        apply (rule image_eqI [where x = \<xi>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
        apply (auto simp: \<open>\<xi> \<in> T\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
        done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
      then obtain \<epsilon> where "\<epsilon>>0" and \<epsilon>: "cball 0 \<epsilon> \<subseteq> U"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
        using \<open>open U\<close> open_contains_cball by blast
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1247
      then have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> cball 0 \<epsilon>"
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1248
                "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> cball 0 \<epsilon>"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
        by (auto simp: norm_mult)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1250
      with \<epsilon> have "\<epsilon> * exp(2 * of_real pi * \<i> * (0/n)) \<in> U"
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1251
                  "\<epsilon> * exp(2 * of_real pi * \<i> * (1/n)) \<in> U" by blast+
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1252
      then obtain y0 y1 where "y0 \<in> T" and y0: "(y0 - \<xi>) * g y0 = \<epsilon> * exp(2 * of_real pi * \<i> * (0/n))"
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1253
                          and "y1 \<in> T" and y1: "(y1 - \<xi>) * g y1 = \<epsilon> * exp(2 * of_real pi * \<i> * (1/n))"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
        by (auto simp: U_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
      then have "y0 \<in> ball \<xi> \<delta>" "y1 \<in> ball \<xi> \<delta>" using Tsb by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
      moreover have "y0 \<noteq> y1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
        using y0 y1 \<open>\<epsilon> > 0\<close> complex_root_unity_eq_1 [of n 1] \<open>n > 0\<close> False by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
      moreover have "T \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
        by (meson Tsb min.cobounded1 order_trans r subset_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
      ultimately have False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
        using inj_onD [OF injf, of y0 y1] \<open>y0 \<in> T\<close> \<open>y1 \<in> T\<close>
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 65538
diff changeset
  1262
        using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
        apply (simp add: y0 y1 power_mult_distrib)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
        apply (force simp: algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
        done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
      then show ?thesis ..
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
text\<open>Hence a nice clean inverse function theorem\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
proposition holomorphic_has_inverse:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
      and "open S" and injf: "inj_on f S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
  obtains g where "g holomorphic_on (f ` S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
                  "\<And>z. z \<in> S \<Longrightarrow> deriv f z * deriv g (f z) = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
                  "\<And>z. z \<in> S \<Longrightarrow> g(f z) = z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
  have ofs: "open (f ` S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
    by (rule open_mapping_thm3 [OF assms])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
  have contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
    by (simp add: holf holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
  have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \<in> S" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
    have 1: "(f has_field_derivative deriv f z) (at z)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1287
      using DERIV_deriv_iff_field_differentiable \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_on_imp_differentiable_at
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
      by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
    have 2: "deriv f z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
      using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
    show ?thesis
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
  1292
      apply (rule has_field_derivative_inverse_strong [OF 1 2 \<open>open S\<close> \<open>z \<in> S\<close>])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
       apply (simp add: holf holomorphic_on_imp_continuous_on)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
      by (simp add: injf the_inv_into_f_f)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
    proof
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
      show "the_inv_into S f holomorphic_on f ` S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
        by (simp add: holomorphic_on_open ofs) (blast intro: *)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
    next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
      fix z assume "z \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
      have "deriv f z \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
        using \<open>z \<in> S\<close> \<open>open S\<close> holf holomorphic_injective_imp_regular injf by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
      then show "deriv f z * deriv (the_inv_into S f) (f z) = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
        using * [OF \<open>z \<in> S\<close>]  by (simp add: DERIV_imp_deriv)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
    next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
      fix z assume "z \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
      show "the_inv_into S f (f z) = z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
        by (simp add: \<open>z \<in> S\<close> injf the_inv_into_f_f)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
subsection\<open>The Schwarz Lemma\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
lemma Schwarz1:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
  assumes holf: "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
      and contf: "continuous_on (closure S) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
      and S: "open S" "connected S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
      and boS: "bounded S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
      and "S \<noteq> {}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
  obtains w where "w \<in> frontier S"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1322
       "\<And>z. z \<in> closure S \<Longrightarrow> norm (f z) \<le> norm (f w)"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
  have connf: "continuous_on (closure S) (norm o f)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
    using contf continuous_on_compose continuous_on_norm_id by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
  have coc: "compact (closure S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
    by (simp add: \<open>bounded S\<close> bounded_closure compact_eq_bounded_closed)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
  then obtain x where x: "x \<in> closure S" and xmax: "\<And>z. z \<in> closure S \<Longrightarrow> norm(f z) \<le> norm(f x)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
    apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
    using \<open>S \<noteq> {}\<close> apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
  then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
  proof (cases "x \<in> frontier S")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
    case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
    then show ?thesis using that xmax by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
    then have "x \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
      using \<open>open S\<close> frontier_def interior_eq x by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
    then have "f constant_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
      apply (rule maximum_modulus_principle [OF holf S \<open>open S\<close> order_refl])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
      using closure_subset apply (blast intro: xmax)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
    then have "f constant_on (closure S)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
      by (rule constant_on_closureI [OF _ contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
    then obtain c where c: "\<And>x. x \<in> closure S \<Longrightarrow> f x = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
      by (meson constant_on_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
    obtain w where "w \<in> frontier S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
      by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
      by (simp add: c frontier_def that)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
lemma Schwarz2:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
 "\<lbrakk>f holomorphic_on ball 0 r;
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
    0 < s; ball w s \<subseteq> ball 0 r;
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
    \<And>z. norm (w-z) < s \<Longrightarrow> norm(f z) \<le> norm(f w)\<rbrakk>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
    \<Longrightarrow> f constant_on ball 0 r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
by (rule maximum_modulus_principle [where U = "ball w s" and \<xi> = w]) (simp_all add: dist_norm)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
lemma Schwarz3:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
  assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
  obtains h where "h holomorphic_on (ball 0 r)" and "\<And>z. norm z < r \<Longrightarrow> f z = z * (h z)" and "deriv f 0 = h 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1366
  define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
  have d0: "deriv f 0 = h 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
    by (simp add: h_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
  moreover have "h holomorphic_on (ball 0 r)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
    by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
  moreover have "norm z < r \<Longrightarrow> f z = z * h z" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
    by (simp add: h_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
  ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
    using that by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
proposition Schwarz_Lemma:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
  assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
      and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
      and \<xi>: "norm \<xi> < 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
    shows "norm (f \<xi>) \<le> norm \<xi>" and "norm(deriv f 0) \<le> 1"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1382
      and "((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z) 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1383
            \<or> norm(deriv f 0) = 1)
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1384
           \<Longrightarrow> \<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1" 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1385
      (is "?P \<Longrightarrow> ?Q")
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
  obtain h where holh: "h holomorphic_on (ball 0 1)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1388
             and fz_eq: "\<And>z. norm z < 1 \<Longrightarrow> f z = z * (h z)" and df0: "deriv f 0 = h 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
    by (rule Schwarz3 [OF holf]) auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
  have noh_le: "norm (h z) \<le> 1" if z: "norm z < 1" for z
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
    have "norm (h z) < a" if a: "1 < a" for a
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
    proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
      have "max (inverse a) (norm z) < 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
        using z a by (simp_all add: inverse_less_1_iff)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
      then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
        using Rats_dense_in_real by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
      then have nzr: "norm z < r" and ira: "inverse r < a"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
        using z a less_imp_inverse_less by force+
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
      then have "0 < r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
        by (meson norm_not_less_zero not_le order.strict_trans2)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
      have holh': "h holomorphic_on ball 0 r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
        by (meson holh \<open>r < 1\<close> holomorphic_on_subset less_eq_real_def subset_ball)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
      have conth': "continuous_on (cball 0 r) h"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
        by (meson \<open>r < 1\<close> dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
      obtain w where w: "norm w = r" and lenw: "\<And>z. norm z < r \<Longrightarrow> norm(h z) \<le> norm(h w)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
        apply (rule Schwarz1 [OF holh']) using conth' \<open>0 < r\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
      have "h w = f w / w" using fz_eq \<open>r < 1\<close> nzr w by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
      then have "cmod (h z) < inverse r"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
        by (metis \<open>0 < r\<close> \<open>r < 1\<close> divide_strict_right_mono inverse_eq_divide
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
                  le_less_trans lenw no norm_divide nzr w)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
      then show ?thesis using ira by linarith
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
    qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
    then show "norm (h z) \<le> 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
      using not_le by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
  show "cmod (f \<xi>) \<le> cmod \<xi>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
  proof (cases "\<xi> = 0")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
    case True then show ?thesis by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
      by (simp add: noh_le fz_eq \<xi> mult_left_le norm_mult)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
  show no_df0: "norm(deriv f 0) \<le> 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
    by (simp add: \<open>\<And>z. cmod z < 1 \<Longrightarrow> cmod (h z) \<le> 1\<close> df0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
  show "?Q" if "?P"
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63152
diff changeset
  1428
    using that
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
  proof
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
    assume "\<exists>z. cmod z < 1 \<and> z \<noteq> 0 \<and> cmod (f z) = cmod z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
    then obtain \<gamma> where \<gamma>: "cmod \<gamma> < 1" "\<gamma> \<noteq> 0" "cmod (f \<gamma>) = cmod \<gamma>" by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
    then have [simp]: "norm (h \<gamma>) = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
      by (simp add: fz_eq norm_mult)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
    have "ball \<gamma> (1 - cmod \<gamma>) \<subseteq> ball 0 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
      by (simp add: ball_subset_ball_iff)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
    moreover have "\<And>z. cmod (\<gamma> - z) < 1 - cmod \<gamma> \<Longrightarrow> cmod (h z) \<le> cmod (h \<gamma>)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
      apply (simp add: algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
      by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
    ultimately obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
      using Schwarz2 [OF holh, of "1 - norm \<gamma>" \<gamma>, unfolded constant_on_def] \<gamma> by auto
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63152
diff changeset
  1441
    then have "norm c = 1"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
      using \<gamma> by force
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63152
diff changeset
  1443
    with c show ?thesis
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
      using fz_eq by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
    assume [simp]: "cmod (deriv f 0) = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
    then obtain c where c: "\<And>z. norm z < 1 \<Longrightarrow> h z = c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
      using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
      by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
    moreover have "norm c = 1"  using df0 c by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
    ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
      using fz_eq by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1456
corollary Schwarz_Lemma':
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1457
  assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0"
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1458
      and no: "\<And>z. norm z < 1 \<Longrightarrow> norm (f z) < 1"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1459
    shows "((\<forall>\<xi>. norm \<xi> < 1 \<longrightarrow> norm (f \<xi>) \<le> norm \<xi>) 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1460
            \<and> norm(deriv f 0) \<le> 1) 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1461
            \<and> (((\<exists>z. norm z < 1 \<and> z \<noteq> 0 \<and> norm(f z) = norm z) 
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1462
              \<or> norm(deriv f 0) = 1)
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1463
              \<longrightarrow> (\<exists>\<alpha>. (\<forall>z. norm z < 1 \<longrightarrow> f z = \<alpha> * z) \<and> norm \<alpha> = 1))"
66793
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1464
  using Schwarz_Lemma [OF assms]
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1465
  by (metis (no_types) norm_eq_zero zero_less_one)
deabce3ccf1f new material about connectedness, etc.
paulson <lp15@cam.ac.uk>
parents: 66660
diff changeset
  1466
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
subsection\<open>The Schwarz reflection principle\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
lemma hol_pal_lem0:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
  assumes "d \<bullet> a \<le> k" "k \<le> d \<bullet> b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
  obtains c where
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
     "c \<in> closed_segment a b" "d \<bullet> c = k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
     "\<And>z. z \<in> closed_segment a c \<Longrightarrow> d \<bullet> z \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
     "\<And>z. z \<in> closed_segment c b \<Longrightarrow> k \<le> d \<bullet> z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
  obtain c where cin: "c \<in> closed_segment a b" and keq: "k = d \<bullet> c"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
    using connected_ivt_hyperplane [of "closed_segment a b" a b d k]
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
    by (auto simp: assms)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
  have "closed_segment a c \<subseteq> {z. d \<bullet> z \<le> k}"  "closed_segment c b \<subseteq> {z. k \<le> d \<bullet> z}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
    unfolding segment_convex_hull using assms keq
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
    by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
  then show ?thesis using cin that by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
lemma hol_pal_lem1:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
  assumes "convex S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
      and abc: "a \<in> S" "b \<in> S" "c \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
          "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k" "d \<bullet> c \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
      and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
      and contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
    shows "contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
           contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
           contour_integral (linepath c a) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
  have "interior (convex hull {a, b, c}) \<subseteq> interior(S \<inter> {x. d \<bullet> x \<le> k})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
    apply (rule interior_mono)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
    apply (rule hull_minimal)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
     apply (simp add: abc lek)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
    apply (rule convex_Int [OF \<open>convex S\<close> convex_halfspace_le])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
  also have "... \<subseteq> {z \<in> S. d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
    by (force simp: interior_open [OF \<open>open S\<close>] \<open>d \<noteq> 0\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
  finally have *: "interior (convex hull {a, b, c}) \<subseteq> {z \<in> S. d \<bullet> z < k}" .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
  have "continuous_on (convex hull {a,b,c}) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
    using \<open>convex S\<close> contf abc continuous_on_subset subset_hull
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
    by fastforce
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
  moreover have "f holomorphic_on interior (convex hull {a,b,c})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
    by (rule holomorphic_on_subset [OF holf1 *])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
  ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
    using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
      by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
lemma hol_pal_lem2:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
  assumes S: "convex S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
      and abc: "a \<in> S" "b \<in> S" "c \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
      and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k" "d \<bullet> b \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
      and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
      and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
      and contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
    shows "contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
           contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
           contour_integral (linepath c a) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
proof (cases "d \<bullet> c \<le> k")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
  case True show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
    by (rule hol_pal_lem1 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
  case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
  then have "d \<bullet> c > k" by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
  obtain a' where a': "a' \<in> closed_segment b c" and "d \<bullet> a' = k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
     and ba': "\<And>z. z \<in> closed_segment b a' \<Longrightarrow> d \<bullet> z \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
     and a'c: "\<And>z. z \<in> closed_segment a' c \<Longrightarrow> k \<le> d \<bullet> z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
    apply (rule hol_pal_lem0 [of d b k c, OF \<open>d \<bullet> b \<le> k\<close>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
    using False by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
  obtain b' where b': "b' \<in> closed_segment a c" and "d \<bullet> b' = k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
     and ab': "\<And>z. z \<in> closed_segment a b' \<Longrightarrow> d \<bullet> z \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
     and b'c: "\<And>z. z \<in> closed_segment b' c \<Longrightarrow> k \<le> d \<bullet> z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
    apply (rule hol_pal_lem0 [of d a k c, OF \<open>d \<bullet> a \<le> k\<close>])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
    using False by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
  have a'b': "a' \<in> S \<and> b' \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
    using a' abc b' convex_contains_segment \<open>convex S\<close> by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
  have "continuous_on (closed_segment c a) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
    by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
  then have 1: "contour_integral (linepath c a) f =
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
                contour_integral (linepath c b') f + contour_integral (linepath b' a) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
    apply (rule contour_integral_split_linepath)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
    using b' by (simp add: closed_segment_commute)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
  have "continuous_on (closed_segment b c) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
    by (meson abc contf continuous_on_subset convex_contains_segment \<open>convex S\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
  then have 2: "contour_integral (linepath b c) f =
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
                contour_integral (linepath b a') f + contour_integral (linepath a' c) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
    by (rule contour_integral_split_linepath [OF _ a'])
62463
547c5c6e66d4 the integral is 0 when otherwise it would be undefined (also for contour integrals)
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
  1553
  have 3: "contour_integral (reversepath (linepath b' a')) f =
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
                - contour_integral (linepath b' a') f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
    by (rule contour_integral_reversepath [OF valid_path_linepath])
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1556
  have fcd_le: "f field_differentiable at x"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
               if "x \<in> interior S \<and> x \<in> interior {x. d \<bullet> x \<le> k}" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1558
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1559
    have "f holomorphic_on S \<inter> {c. d \<bullet> c < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1560
      by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
    then have "\<exists>C D. x \<in> interior C \<inter> interior D \<and> f holomorphic_on interior C \<inter> interior D"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
      using that
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
      by (metis Collect_mem_eq Int_Collect \<open>d \<noteq> 0\<close> interior_halfspace_le interior_open \<open>open S\<close>)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1564
    then show "f field_differentiable at x"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1565
      by (metis at_within_interior holomorphic_on_def interior_Int interior_interior)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1566
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1567
  have ab_le: "\<And>x. x \<in> closed_segment a b \<Longrightarrow> d \<bullet> x \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1568
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1569
    fix x :: complex
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1570
    assume "x \<in> closed_segment a b"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
    then have "\<And>C. x \<in> C \<or> b \<notin> C \<or> a \<notin> C \<or> \<not> convex C"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
      by (meson contra_subsetD convex_contains_segment)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
    then show "d \<bullet> x \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
      by (metis lek convex_halfspace_le mem_Collect_eq)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
  have "continuous_on (S \<inter> {x. d \<bullet> x \<le> k}) f" using contf
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
    by (simp add: continuous_on_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
  then have "(f has_contour_integral 0)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
         (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)"
68310
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1580
    apply (rule Cauchy_theorem_convex [where K = "{}"])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
    apply (simp_all add: path_image_join convex_Int convex_halfspace_le \<open>convex S\<close> fcd_le ab_le
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
                closed_segment_subset abc a'b' ba')
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
    by (metis \<open>d \<bullet> a' = k\<close> \<open>d \<bullet> b' = k\<close> convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
  then have 4: "contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
                contour_integral (linepath b a') f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
                contour_integral (linepath a' b') f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
                contour_integral (linepath b' a) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
    by (rule has_chain_integral_chain_integral4)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1589
  have fcd_ge: "f field_differentiable at x"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1590
               if "x \<in> interior S \<and> x \<in> interior {x. k \<le> d \<bullet> x}" for x
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1591
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
    have f2: "f holomorphic_on S \<inter> {c. k < d \<bullet> c}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
      by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
    have f3: "interior S = S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
      by (simp add: interior_open \<open>open S\<close>)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
    then have "x \<in> S \<inter> interior {c. k \<le> d \<bullet> c}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
      using that by simp
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1598
    then show "f field_differentiable at x"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
      using f3 f2 unfolding holomorphic_on_def
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
      by (metis (no_types) \<open>d \<noteq> 0\<close> at_within_interior interior_Int interior_halfspace_ge interior_interior)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
  have "continuous_on (S \<inter> {x. k \<le> d \<bullet> x}) f" using contf
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
    by (simp add: continuous_on_subset)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
  then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')"
68310
d0a7ddf5450e more general tidying
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1605
    apply (rule Cauchy_theorem_convex [where K = "{}"])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
    apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \<open>convex S\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
                      fcd_ge closed_segment_subset abc a'b' a'c)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
    by (metis \<open>d \<bullet> a' = k\<close> b'c closed_segment_commute convex_contains_segment
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
              convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
  then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
    by (rule has_chain_integral_chain_integral3)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
    using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
lemma hol_pal_lem3:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
  assumes S: "convex S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
      and abc: "a \<in> S" "b \<in> S" "c \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
      and "d \<noteq> 0" and lek: "d \<bullet> a \<le> k"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
      and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
      and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
      and contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
    shows "contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
           contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
           contour_integral (linepath c a) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
proof (cases "d \<bullet> b \<le> k")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
  case True show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
    by (rule hol_pal_lem2 [OF S abc \<open>d \<noteq> 0\<close> lek True holf1 holf2 contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
  case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
  proof (cases "d \<bullet> c \<le> k")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
    case True
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
    have "contour_integral (linepath c a) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
          contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
          contour_integral (linepath b c) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
      by (rule hol_pal_lem2 [OF S \<open>c \<in> S\<close> \<open>a \<in> S\<close> \<open>b \<in> S\<close> \<open>d \<noteq> 0\<close> \<open>d \<bullet> c \<le> k\<close> lek holf1 holf2 contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
      by (simp add: algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
  next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
    case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
    have "contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
          contour_integral (linepath c a) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
          contour_integral (linepath a b) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
      apply (rule hol_pal_lem2 [OF S \<open>b \<in> S\<close> \<open>c \<in> S\<close> \<open>a \<in> S\<close>, of "-d" "-k"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
      using \<open>d \<noteq> 0\<close> \<open>\<not> d \<bullet> b \<le> k\<close> False by (simp_all add: holf1 holf2 contf)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
    then show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
      by (simp add: algebra_simps)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
lemma hol_pal_lem4:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
  assumes S: "convex S" "open S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
      and abc: "a \<in> S" "b \<in> S" "c \<in> S" and "d \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
      and holf1: "f holomorphic_on {z. z \<in> S \<and> d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
      and holf2: "f holomorphic_on {z. z \<in> S \<and> k < d \<bullet> z}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
      and contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
    shows "contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
           contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
           contour_integral (linepath c a) f = 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
proof (cases "d \<bullet> a \<le> k")
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
  case True show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
    by (rule hol_pal_lem3 [OF S abc \<open>d \<noteq> 0\<close> True holf1 holf2 contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
next
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
  case False
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
    apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
    using \<open>d \<noteq> 0\<close> False by (simp_all add: holf1 holf2 contf)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1671
lemma holomorphic_on_paste_across_line:
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
  assumes S: "open S" and "d \<noteq> 0"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
      and holf1: "f holomorphic_on (S \<inter> {z. d \<bullet> z < k})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1674
      and holf2: "f holomorphic_on (S \<inter> {z. k < d \<bullet> z})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
      and contf: "continuous_on S f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
    shows "f holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
  have *: "\<exists>t. open t \<and> p \<in> t \<and> continuous_on t f \<and>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
               (\<forall>a b c. convex hull {a, b, c} \<subseteq> t \<longrightarrow>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
                         contour_integral (linepath a b) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
                         contour_integral (linepath b c) f +
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
                         contour_integral (linepath c a) f = 0)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
          if "p \<in> S" for p
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
  proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
    obtain e where "e>0" and e: "ball p e \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
      using \<open>p \<in> S\<close> openE S by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
    then have "continuous_on (ball p e) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
      using contf continuous_on_subset by blast
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
    moreover have "f holomorphic_on {z. dist p z < e \<and> d \<bullet> z < k}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
      apply (rule holomorphic_on_subset [OF holf1])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
      using e by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
    moreover have "f holomorphic_on {z. dist p z < e \<and> k < d \<bullet> z}"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
      apply (rule holomorphic_on_subset [OF holf2])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
      using e by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
    ultimately show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
      apply (rule_tac x="ball p e" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
      using \<open>e > 0\<close> e \<open>d \<noteq> 0\<close>
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
      apply (simp add:, clarify)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
      apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
      apply (auto simp: subset_hull)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
      done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
  qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
  show ?thesis
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
    by (blast intro: * Morera_local_triangle analytic_imp_holomorphic)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
proposition Schwarz_reflection:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
  assumes "open S" and cnjs: "cnj ` S \<subseteq> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
      and  holf: "f holomorphic_on (S \<inter> {z. 0 < Im z})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
      and contf: "continuous_on (S \<inter> {z. 0 \<le> Im z}) f"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
      and f: "\<And>z. \<lbrakk>z \<in> S; z \<in> \<real>\<rbrakk> \<Longrightarrow> (f z) \<in> \<real>"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
    shows "(\<lambda>z. if 0 \<le> Im z then f z else cnj(f(cnj z))) holomorphic_on S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
  have 1: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. 0 < Im z})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
    by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
  have cont_cfc: "continuous_on (S \<inter> {z. Im z \<le> 0}) (cnj o f o cnj)"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
    apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
    using cnjs apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
    done
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1720
  have "cnj \<circ> f \<circ> cnj field_differentiable at x within S \<inter> {z. Im z < 0}"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1721
        if "x \<in> S" "Im x < 0" "f field_differentiable at (cnj x) within S \<inter> {z. 0 < Im z}" for x
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
    using that
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
  1723
    apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify)
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
    apply (rule_tac x="cnj f'" in exI)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
    apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
    apply (drule_tac x="cnj xa" in bspec)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
    using cnjs apply force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
    apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
  then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \<inter> {z. Im z < 0})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
    using holf cnjs
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
    by (force simp: holomorphic_on_def)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
  have 2: "(\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z))) holomorphic_on (S \<inter> {z. Im z < 0})"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
    apply (rule iffD1 [OF holomorphic_cong [OF refl]])
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
    using hol_cfc by auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
  have [simp]: "(S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}) = S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
    by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
  have "continuous_on ((S \<inter> {z. 0 \<le> Im z}) \<union> (S \<inter> {z. Im z \<le> 0}))
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
                       (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
    apply (rule continuous_on_cases_local)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
    using cont_cfc contf
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
    apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge)
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
    using f Reals_cnj_iff complex_is_Real_iff apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
  then have 3: "continuous_on S (\<lambda>z. if 0 \<le> Im z then f z else cnj (f (cnj z)))"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
    by force
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
  show ?thesis
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  1748
    apply (rule holomorphic_on_paste_across_line [OF \<open>open S\<close>, of "- \<i>" _ 0])
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
    using 1 2 3
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
    apply auto
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
    done
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1754
subsection\<open>Bloch's theorem\<close>
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1755
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1756
lemma Bloch_lemma_0:
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1757
  assumes holf: "f holomorphic_on cball 0 r" and "0 < r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1758
      and [simp]: "f 0 = 0"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1759
      and le: "\<And>z. norm z < r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f 0)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1760
    shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \<subseteq> f ` ball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1761
proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1762
  have "sqrt 2 < 3/2"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1763
    by (rule real_less_lsqrt) (auto simp: power2_eq_square)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1764
  then have sq3: "0 < 3 - 2 * sqrt 2" by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1765
  show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1766
  proof (cases "deriv f 0 = 0")
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1767
    case True then show ?thesis by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1768
  next
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1769
    case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1770
    define C where "C = 2 * norm(deriv f 0)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1771
    have "0 < C" using False by (simp add: C_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1772
    have holf': "f holomorphic_on ball 0 r" using holf
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1773
      using ball_subset_cball holomorphic_on_subset by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1774
    then have holdf': "deriv f holomorphic_on ball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1775
      by (rule holomorphic_deriv [OF _ open_ball])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1776
    have "Le1": "norm(deriv f z - deriv f 0) \<le> norm z / (r - norm z) * C"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1777
                if "norm z < r" for z
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1778
    proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1779
      have T1: "norm(deriv f z - deriv f 0) \<le> norm z / (R - norm z) * C"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1780
              if R: "norm z < R" "R < r" for R
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1781
      proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1782
        have "0 < R" using R
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1783
          by (metis less_trans norm_zero zero_less_norm_iff)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1784
        have df_le: "\<And>x. norm x < r \<Longrightarrow> norm (deriv f x) \<le> C"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1785
          using le by (simp add: C_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1786
        have hol_df: "deriv f holomorphic_on cball 0 R"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1787
          apply (rule holomorphic_on_subset) using R holdf' by auto
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1788
        have *: "((\<lambda>w. deriv f w / (w - z)) has_contour_integral 2 * pi * \<i> * deriv f z) (circlepath 0 R)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1789
                 if "norm z < R" for z
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1790
          using \<open>0 < R\<close> that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"]
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1791
          by (force simp: winding_number_circlepath)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1792
        have **: "((\<lambda>x. deriv f x / (x - z) - deriv f x / x) has_contour_integral
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1793
                   of_real (2 * pi) * \<i> * (deriv f z - deriv f 0))
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1794
                  (circlepath 0 R)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1795
           using has_contour_integral_diff [OF * [of z] * [of 0]] \<open>0 < R\<close> that
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1796
           by (simp add: algebra_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1797
        have [simp]: "\<And>x. norm x = R \<Longrightarrow> x \<noteq> z"  using that(1) by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1798
        have "norm (deriv f x / (x - z) - deriv f x / x)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1799
                     \<le> C * norm z / (R * (R - norm z))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1800
                  if "norm x = R" for x
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1801
        proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1802
          have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) =
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1803
                        norm (deriv f x) * norm z"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1804
            by (simp add: norm_mult right_diff_distrib')
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1805
          show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1806
            using  \<open>0 < R\<close> \<open>0 < C\<close> R that
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1807
            apply (simp add: norm_mult norm_divide divide_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1808
            using df_le norm_triangle_ineq2 \<open>0 < C\<close> apply (auto intro!: mult_mono)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1809
            done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1810
        qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1811
        then show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1812
          using has_contour_integral_bound_circlepath
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1813
                  [OF **, of "C * norm z/(R*(R - norm z))"]
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1814
                \<open>0 < R\<close> \<open>0 < C\<close> R
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1815
          apply (simp add: norm_mult norm_divide)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1816
          apply (simp add: divide_simps mult.commute)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1817
          done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1818
      qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1819
      obtain r' where r': "norm z < r'" "r' < r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1820
        using Rats_dense_in_real [of "norm z" r] \<open>norm z < r\<close> by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1821
      then have [simp]: "closure {r'<..<r} = {r'..r}" by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1822
      show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1823
        apply (rule continuous_ge_on_closure
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1824
                 [where f = "\<lambda>r. norm z / (r - norm z) * C" and s = "{r'<..<r}",
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1825
                  OF _ _ T1])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1826
        apply (intro continuous_intros)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1827
        using that r'
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1828
        apply (auto simp: not_le)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1829
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1830
    qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1831
    have "*": "(norm z - norm z^2/(r - norm z)) * norm(deriv f 0) \<le> norm(f z)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1832
              if r: "norm z < r" for z
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1833
    proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1834
      have 1: "\<And>x. x \<in> ball 0 r \<Longrightarrow>
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1835
              ((\<lambda>z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1836
               (at x within ball 0 r)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1837
        by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1838
      have 2: "closed_segment 0 z \<subseteq> ball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1839
        by (metis \<open>0 < r\<close> convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1840
      have 3: "(\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1841
        apply (rule integrable_on_cmult_right [where 'b=real, simplified])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1842
        apply (rule integrable_on_cdivide [where 'b=real, simplified])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1843
        apply (rule integrable_on_cmult_left [where 'b=real, simplified])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1844
        apply (rule ident_integrable_on)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1845
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1846
      have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm z * norm z * x * C / (r - norm z)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1847
              if x: "0 \<le> x" "x \<le> 1" for x
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1848
      proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1849
        have [simp]: "x * norm z < r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1850
          using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1851
        have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \<le> norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1852
          apply (rule Le1) using r x \<open>0 < r\<close> by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1853
        also have "... \<le> norm (x *\<^sub>R z) / (r - norm z) * C"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1854
          using r x \<open>0 < r\<close>
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1855
          apply (simp add: divide_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1856
          by (simp add: \<open>0 < C\<close> mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1857
        finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \<le> norm (x *\<^sub>R z)  / (r - norm z) * C * norm z"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1858
          by (rule mult_right_mono) simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1859
        with x show ?thesis by (simp add: algebra_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1860
      qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1861
      have le_norm: "abc \<le> norm d - e \<Longrightarrow> norm(f - d) \<le> e \<Longrightarrow> abc \<le> norm f" for abc d e and f::complex
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1862
        by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1863
      have "norm (integral {0..1} (\<lambda>x. (deriv f (x *\<^sub>R z) - deriv f 0) * z))
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1864
            \<le> integral {0..1} (\<lambda>t. (norm z)\<^sup>2 * t / (r - norm z) * C)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1865
        apply (rule integral_norm_bound_integral)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1866
        using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1867
        apply (simp add: has_contour_integral_linepath has_integral_integrable_integral)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1868
        apply (rule 3)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1869
        apply (simp add: norm_mult power2_eq_square 4)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1870
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1871
      then have int_le: "norm (f z - deriv f 0 * z) \<le> (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1872
        using contour_integral_primitive [OF 1, of "linepath 0 z"] 2
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1873
        apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1874
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1875
      show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1876
        apply (rule le_norm [OF _ int_le])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1877
        using \<open>norm z < r\<close>
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1878
        apply (simp add: power2_eq_square divide_simps C_def norm_mult)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1879
        proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1880
          have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \<le> norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)"
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 70346
diff changeset
  1881
            by (simp add: algebra_simps)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1882
          then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \<le> norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)"
70347
e5cd5471c18a official fact collection sign_simps
haftmann
parents: 70346
diff changeset
  1883
            by (simp add: algebra_simps)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1884
        qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1885
    qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1886
    have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2)  < 1"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1887
      by (auto simp:  sqrt2_less_2)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1888
    have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1889
      apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1890
      apply (subst closure_ball)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1891
      using \<open>0 < r\<close> mult_pos_pos sq201
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1892
      apply (auto simp: cball_subset_cball_iff)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1893
      done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1894
    have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1895
      apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1896
      using \<open>0 < r\<close> mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1897
      using False \<open>0 < r\<close> centre_in_ball holf' holomorphic_nonconstant by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1898
    have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) =
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1899
          ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1900
      by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1901
    also have "...  \<subseteq> f ` ball 0 ((1 - sqrt 2 / 2) * r)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1902
    proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1903
      have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \<le> norm (f z)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1904
           if "norm z = (1 - sqrt 2 / 2) * r" for z
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1905
        apply (rule order_trans [OF _ *])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1906
        using  \<open>0 < r\<close>
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1907
        apply (simp_all add: field_simps  power2_eq_square that)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1908
        apply (simp add: mult.assoc [symmetric])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1909
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1910
      show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1911
        apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1912
        using \<open>0 < r\<close> sq201 3 apply simp_all
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1913
        using C_def \<open>0 < C\<close> sq3 apply force
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1914
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1915
     qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1916
    also have "...  \<subseteq> f ` ball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1917
      apply (rule image_subsetI [OF imageI], simp)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1918
      apply (erule less_le_trans)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1919
      using \<open>0 < r\<close> apply (auto simp: field_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1920
      done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1921
    finally show ?thesis .
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1922
  qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1923
qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1924
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1925
lemma Bloch_lemma:
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1926
  assumes holf: "f holomorphic_on cball a r" and "0 < r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1927
      and le: "\<And>z. z \<in> ball a r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f a)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1928
    shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \<subseteq> f ` ball a r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1929
proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1930
  have fz: "(\<lambda>z. f (a + z)) = f o (\<lambda>z. (a + z))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1931
    by (simp add: o_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1932
  have hol0: "(\<lambda>z. f (a + z)) holomorphic_on cball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1933
    unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1934
  then have [simp]: "\<And>x. norm x < r \<Longrightarrow> (\<lambda>z. f (a + z)) field_differentiable at x"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  1935
    by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1936
  have [simp]: "\<And>z. norm z < r \<Longrightarrow> f field_differentiable at (a + z)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1937
    by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1938
  then have [simp]: "f field_differentiable at a"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1939
    by (metis add.comm_neutral \<open>0 < r\<close> norm_eq_zero)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1940
  have hol1: "(\<lambda>z. f (a + z) - f a) holomorphic_on cball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1941
    by (intro holomorphic_intros hol0)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1942
  then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\<lambda>z. f (a + z) - f a) 0))
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1943
             \<subseteq> (\<lambda>z. f (a + z) - f a) ` ball 0 r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1944
    apply (rule Bloch_lemma_0)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1945
    apply (simp_all add: \<open>0 < r\<close>)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1946
    apply (simp add: fz complex_derivative_chain)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1947
    apply (simp add: dist_norm le)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1948
    done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1949
  then show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1950
    apply clarify
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1951
    apply (drule_tac c="x - f a" in subsetD)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1952
     apply (force simp: fz \<open>0 < r\<close> dist_norm complex_derivative_chain field_differentiable_compose)+
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1953
    done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1954
qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1955
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1956
proposition Bloch_unit:
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1957
  assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  1958
  obtains b r where "1/12 < r" and "ball b r \<subseteq> f ` (ball a 1)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1959
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1960
  define r :: real where "r = 249/256"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1961
  have "0 < r" "r < 1" by (auto simp: r_def)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1962
  define g where "g z = deriv f z * of_real(r - norm(z - a))" for z
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1963
  have "deriv f holomorphic_on ball a 1"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1964
    by (rule holomorphic_deriv [OF holf open_ball])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1965
  then have "continuous_on (ball a 1) (deriv f)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1966
    using holomorphic_on_imp_continuous_on by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1967
  then have "continuous_on (cball a r) (deriv f)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1968
    by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \<open>r < 1\<close>)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1969
  then have "continuous_on (cball a r) g"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1970
    by (simp add: g_def continuous_intros)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1971
  then have 1: "compact (g ` cball a r)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1972
    by (rule compact_continuous_image [OF _ compact_cball])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1973
  have 2: "g ` cball a r \<noteq> {}"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1974
    using \<open>r > 0\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1975
  obtain p where pr: "p \<in> cball a r"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1976
             and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1977
    using distance_attains_sup [OF 1 2, of 0] by force
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1978
  define t where "t = (r - norm(p - a)) / 2"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1979
  have "norm (p - a) \<noteq> r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1980
    using pge [of a] \<open>r > 0\<close> by (auto simp: g_def norm_mult)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1981
  then have "norm (p - a) < r" using pr
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1982
    by (simp add: norm_minus_commute dist_norm)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1983
  then have "0 < t"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1984
    by (simp add: t_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1985
  have cpt: "cball p t \<subseteq> ball a r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1986
    using \<open>0 < t\<close> by (simp add: cball_subset_ball_iff dist_norm t_def field_simps)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1987
  have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1988
            if "y \<in> cball a r" for y
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1989
  proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1990
    have [simp]: "norm (y - a) \<le> r"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1991
      using that by (simp add: dist_norm norm_minus_commute)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1992
    have "norm (g y) \<le> norm (g p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1993
      using pge [OF that] by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1994
    then have "norm (deriv f y) * abs (r - norm (y - a)) \<le> norm (deriv f p) * abs (r - norm (p - a))"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1995
      by (simp only: dist_norm g_def norm_mult norm_of_real)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1996
    with that \<open>norm (p - a) < r\<close> show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1997
      by (simp add: dist_norm divide_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1998
  qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  1999
  have le_norm_dfp: "r / (r - norm (p - a)) \<le> norm (deriv f p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2000
    using gen_le_dfp [of a] \<open>r > 0\<close> by auto
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2001
  have 1: "f holomorphic_on cball p t"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2002
    apply (rule holomorphic_on_subset [OF holf])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2003
    using cpt \<open>r < 1\<close> order_subst1 subset_ball by auto
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2004
  have 2: "norm (deriv f z) \<le> 2 * norm (deriv f p)" if "z \<in> ball p t" for z
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2005
  proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2006
    have z: "z \<in> cball a r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2007
      by (meson ball_subset_cball subsetD cpt that)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2008
    then have "norm(z - a) < r"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2009
      by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2010
    have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2011
      using gen_le_dfp [OF z] by simp
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2012
    with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close>
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2013
    have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2014
       by (simp add: field_simps)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2015
    also have "... \<le> 2 * norm (deriv f p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2016
      apply (rule mult_right_mono)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2017
      using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close>
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2018
      apply (simp_all add: field_simps t_def dist_norm [symmetric])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2019
      using dist_triangle3 [of z a p] by linarith
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2020
    finally show ?thesis .
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2021
  qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2022
  have sqrt2: "sqrt 2 < 2113/1494"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2023
    by (rule real_less_lsqrt) (auto simp: power2_eq_square)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2024
  then have sq3: "0 < 3 - 2 * sqrt 2" by simp
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2025
  have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2026
    using sq3 sqrt2 by (auto simp: field_simps r_def)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2027
  also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2028
    using \<open>norm (p - a) < r\<close> le_norm_dfp   by (simp add: pos_divide_le_eq)
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2029
  finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2030
    using pos_divide_less_eq half_gt_zero_iff sq3 by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2031
  then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2032
    using sq3 by (simp add: mult.commute t_def)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2033
  have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball p t"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2034
    by (rule Bloch_lemma [OF 1 \<open>0 < t\<close> 2])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2035
  also have "... \<subseteq> f ` ball a 1"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2036
    apply (rule image_mono)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2037
    apply (rule order_trans [OF ball_subset_cball])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2038
    apply (rule order_trans [OF cpt])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2039
    using \<open>0 < t\<close> \<open>r < 1\<close> apply (simp add: ball_subset_ball_iff dist_norm)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2040
    done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2041
  finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \<subseteq> f ` ball a 1" .
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2042
  with ** show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2043
    by (rule that)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2044
qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2045
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2046
theorem Bloch:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2047
  assumes holf: "f holomorphic_on ball a r" and "0 < r"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2048
      and r': "r' \<le> r * norm (deriv f a) / 12"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2049
  obtains b where "ball b r' \<subseteq> f ` (ball a r)"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2050
proof (cases "deriv f a = 0")
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2051
  case True with r' show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2052
    using ball_eq_empty that by fastforce
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2053
next
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2054
  case False
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2055
  define C where "C = deriv f a"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2056
  have "0 < norm C" using False by (simp add: C_def)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2057
  have dfa: "f field_differentiable at a"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2058
    apply (rule holomorphic_on_imp_differentiable_at [OF holf])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2059
    using \<open>0 < r\<close> by auto
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2060
  have fo: "(\<lambda>z. f (a + of_real r * z)) = f o (\<lambda>z. (a + of_real r * z))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2061
    by (simp add: o_def)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2062
  have holf': "f holomorphic_on (\<lambda>z. a + complex_of_real r * z) ` ball 0 1"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2063
    apply (rule holomorphic_on_subset [OF holf])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2064
    using \<open>0 < r\<close> apply (force simp: dist_norm norm_mult)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2065
    done
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2066
  have 1: "(\<lambda>z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2067
    apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2068
    using \<open>0 < r\<close> by (simp add: C_def False)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2069
  have "((\<lambda>z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2070
        (deriv f (a + of_real r * z) / C)) (at z)"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2071
       if "norm z < 1" for z
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2072
  proof -
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2073
    have *: "((\<lambda>x. f (a + of_real r * x)) has_field_derivative
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2074
           (deriv f (a + of_real r * z) * of_real r)) (at z)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2075
      apply (simp add: fo)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2076
      apply (rule DERIV_chain [OF field_differentiable_derivI])
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2077
      apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2078
      using \<open>0 < r\<close> apply (simp add: dist_norm norm_mult that)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2079
      apply (rule derivative_eq_intros | simp)+
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2080
      done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2081
    show ?thesis
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2082
      apply (rule derivative_eq_intros * | simp)+
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2083
      using \<open>0 < r\<close> by (auto simp: C_def False)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2084
  qed
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2085
  have 2: "deriv (\<lambda>z. f (a + of_real r * z) / (C * of_real r)) 0 = 1"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2086
    apply (subst deriv_cdivide_right)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2087
    apply (simp add: field_differentiable_def fo)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2088
    apply (rule exI)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2089
    apply (rule DERIV_chain [OF field_differentiable_derivI])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2090
    apply (simp add: dfa)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2091
    apply (rule derivative_eq_intros | simp add: C_def False fo)+
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2092
    using \<open>0 < r\<close>
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2093
    apply (simp add: C_def False fo)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2094
    apply (simp add: derivative_intros dfa complex_derivative_chain)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2095
    done
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68629
diff changeset
  2096
  have sb1: "(*) (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2097
             \<subseteq> f ` ball a r"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2098
    using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68629
diff changeset
  2099
  have sb2: "ball (C * r * b) r' \<subseteq> (*) (C * r) ` ball b t"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2100
             if "1 / 12 < t" for b t
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2101
  proof -
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2102
    have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2103
      using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2104
      by auto
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2105
    show ?thesis
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2106
      apply clarify
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2107
      apply (rule_tac x="x / (C * r)" in image_eqI)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2108
      using \<open>0 < r\<close>
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2109
      apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2110
      apply (erule less_le_trans)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2111
      apply (rule order_trans [OF r' *])
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2112
      done
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2113
  qed
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2114
  show ?thesis
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2115
    apply (rule Bloch_unit [OF 1 2])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2116
    apply (rename_tac t)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2117
    apply (rule_tac b="(C * of_real r) * b" in that)
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2118
    apply (drule image_mono [where f = "\<lambda>z. (C * of_real r) * z"])
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2119
    using sb1 sb2
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2120
    apply force
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2121
    done
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2122
qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2123
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2124
corollary Bloch_general:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2125
  assumes holf: "f holomorphic_on s" and "a \<in> s"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2126
      and tle: "\<And>z. z \<in> frontier s \<Longrightarrow> t \<le> dist a z"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2127
      and rle: "r \<le> t * norm(deriv f a) / 12"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2128
  obtains b where "ball b r \<subseteq> f ` s"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2129
proof -
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2130
  consider "r \<le> 0" | "0 < t * norm(deriv f a) / 12" using rle by force
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2131
  then show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2132
  proof cases
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2133
    case 1 then show ?thesis
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  2134
      by (simp add: ball_empty that)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2135
  next
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2136
    case 2
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2137
    show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2138
    proof (cases "deriv f a = 0")
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2139
      case True then show ?thesis
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  2140
        using rle by (simp add: ball_empty that)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2141
    next
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2142
      case False
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2143
      then have "t > 0"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2144
        using 2 by (force simp: zero_less_mult_iff)
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69064
diff changeset
  2145
      have "\<not> ball a t \<subseteq> s \<Longrightarrow> ball a t \<inter> frontier s \<noteq> {}"
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2146
        apply (rule connected_Int_frontier [of "ball a t" s], simp_all)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2147
        using \<open>0 < t\<close> \<open>a \<in> s\<close> centre_in_ball apply blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2148
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2149
      with tle have *: "ball a t \<subseteq> s" by fastforce
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2150
      then have 1: "f holomorphic_on ball a t"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2151
        using holf using holomorphic_on_subset by blast
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2152
      show ?thesis
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2153
        apply (rule Bloch [OF 1 \<open>t > 0\<close> rle])
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2154
        apply (rule_tac b=b in that)
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2155
        using * apply force
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2156
        done
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2157
    qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2158
  qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2159
qed
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62463
diff changeset
  2160
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  2161
subsection \<open>Cauchy's residue theorem\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2162
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2163
text\<open>Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem.
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2164
    Interactive Theorem Proving\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2165
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
  2166
definition\<^marker>\<open>tag important\<close> residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2167
  "residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  2168
    \<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2169
66286
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2170
lemma Eps_cong:
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2171
  assumes "\<And>x. P x = Q x"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2172
  shows   "Eps P = Eps Q"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2173
  using ext[of P Q, OF assms] by simp
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2174
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2175
lemma residue_cong:
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2176
  assumes eq: "eventually (\<lambda>z. f z = g z) (at z)" and "z = z'"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2177
  shows   "residue f z = residue g z'"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2178
proof -
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2179
  from assms have eq': "eventually (\<lambda>z. g z = f z) (at z)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2180
    by (simp add: eq_commute)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2181
  let ?P = "\<lambda>f c e. (\<forall>\<epsilon>>0. \<epsilon> < e \<longrightarrow>
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2182
   (f has_contour_integral of_real (2 * pi) * \<i> * c) (circlepath z \<epsilon>))"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2183
  have "residue f z = residue g z" unfolding residue_def
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2184
  proof (rule Eps_cong)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2185
    fix c :: complex
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2186
    have "\<exists>e>0. ?P g c e" 
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2187
      if "\<exists>e>0. ?P f c e" and "eventually (\<lambda>z. f z = g z) (at z)" for f g 
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2188
    proof -
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2189
      from that(1) obtain e where e: "e > 0" "?P f c e"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2190
        by blast
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2191
      from that(2) obtain e' where e': "e' > 0" "\<And>z'. z' \<noteq> z \<Longrightarrow> dist z' z < e' \<Longrightarrow> f z' = g z'"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2192
        unfolding eventually_at by blast
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2193
      have "?P g c (min e e')"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2194
      proof (intro allI exI impI, goal_cases)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2195
        case (1 \<epsilon>)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2196
        hence "(f has_contour_integral of_real (2 * pi) * \<i> * c) (circlepath z \<epsilon>)" 
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2197
          using e(2) by auto
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2198
        thus ?case
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2199
        proof (rule has_contour_integral_eq)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2200
          fix z' assume "z' \<in> path_image (circlepath z \<epsilon>)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2201
          hence "dist z' z < e'" and "z' \<noteq> z"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2202
            using 1 by (auto simp: dist_commute)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2203
          with e'(2)[of z'] show "f z' = g z'" by simp
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2204
        qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2205
      qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2206
      moreover from e and e' have "min e e' > 0" by auto
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2207
      ultimately show ?thesis by blast
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2208
    qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2209
    from this[OF _ eq] and this[OF _ eq']
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2210
      show "(\<exists>e>0. ?P f c e) \<longleftrightarrow> (\<exists>e>0. ?P g c e)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2211
      by blast
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2212
  qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2213
  with assms show ?thesis by simp
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2214
qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2215
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2216
lemma contour_integral_circlepath_eq:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2217
  assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2218
    and e2_cball:"cball z e2 \<subseteq> s"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2219
  shows
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2220
    "f contour_integrable_on circlepath z e1"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2221
    "f contour_integrable_on circlepath z e2"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2222
    "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2223
proof -
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2224
  define l where "l \<equiv> linepath (z+e2) (z+e1)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2225
  have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2226
  have "e2>0" using \<open>e1>0\<close> \<open>e1\<le>e2\<close> by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2227
  have zl_img:"z\<notin>path_image l"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2228
    proof
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2229
      assume "z \<in> path_image l"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2230
      then have "e2 \<le> cmod (e2 - e1)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2231
        using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \<open>e1>0\<close> \<open>e2>0\<close> unfolding l_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2232
        by (auto simp add:closed_segment_commute)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2233
      thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2234
        apply (subst (asm) norm_of_real)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2235
        by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2236
    qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2237
  define g where "g \<equiv> circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2238
  show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2239
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2240
      show "f contour_integrable_on circlepath z e2"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2241
        apply (intro contour_integrable_continuous_circlepath[OF
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2242
                continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2243
        using \<open>e2>0\<close> e2_cball by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2244
      show "f contour_integrable_on (circlepath z e1)"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2245
        apply (intro contour_integrable_continuous_circlepath[OF
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2246
                      continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2247
        using \<open>e1>0\<close> \<open>e1\<le>e2\<close> e2_cball by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2248
    qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2249
  have [simp]:"f contour_integrable_on l"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2250
    proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2251
      have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2252
        by (intro closed_segment_subset,auto simp add:dist_norm)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2253
      hence "closed_segment (z + e2) (z + e1) \<subseteq> s - {z}" using zl_img e2_cball unfolding l_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2254
        by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2255
      then show "f contour_integrable_on l" unfolding l_def
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2256
        apply (intro contour_integrable_continuous_linepath[OF
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2257
                      continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2258
        by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2259
    qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2260
  let ?ig="\<lambda>g. contour_integral g f"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2261
  have "(f has_contour_integral 0) g"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2262
    proof (rule Cauchy_theorem_global[OF _ f_holo])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2263
      show "open (s - {z})" using \<open>open s\<close> by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2264
      show "valid_path g" unfolding g_def l_def by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2265
      show "pathfinish g = pathstart g" unfolding g_def l_def by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2266
    next
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2267
      have path_img:"path_image g \<subseteq> cball z e2"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2268
        proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2269
          have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2270
            by (intro closed_segment_subset,auto simp add:dist_norm)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2271
          moreover have "sphere z \<bar>e1\<bar> \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1\<le>e2\<close> \<open>e1>0\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2272
          ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2273
            by (simp add: path_image_join closed_segment_commute)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2274
        qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2275
      show "path_image g \<subseteq> s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2276
        proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2277
          have "z\<notin>path_image g" using zl_img
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2278
            unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2279
          moreover note \<open>cball z e2 \<subseteq> s\<close> and path_img
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2280
          ultimately show ?thesis by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2281
        qed
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2282
      show "winding_number g w = 0" when"w \<notin> s - {z}" for w
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2283
        proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2284
          have "winding_number g w = 0" when "w\<notin>s" using that e2_cball
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2285
            apply (intro winding_number_zero_outside[OF _ _ _ _ path_img])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2286
            by (auto simp add:g_def l_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2287
          moreover have "winding_number g z=0"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2288
            proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2289
              let ?Wz="\<lambda>g. winding_number g z"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2290
              have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1))
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2291
                  + ?Wz (reversepath l)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2292
                using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2293
                by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2294
              also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2295
                using zl_img
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2296
                apply (subst (2) winding_number_reversepath)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2297
                by (auto simp add:l_def closed_segment_commute)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2298
              also have "... = 0"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2299
                proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2300
                  have "?Wz (circlepath z e2) = 1" using \<open>e2>0\<close>
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2301
                    by (auto intro: winding_number_circlepath_centre)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2302
                  moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \<open>e1>0\<close>
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2303
                    apply (subst winding_number_reversepath)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2304
                    by (auto intro: winding_number_circlepath_centre)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2305
                  ultimately show ?thesis by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2306
                qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2307
              finally show ?thesis .
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2308
            qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2309
          ultimately show ?thesis using that by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2310
        qed
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2311
    qed
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2312
  then have "0 = ?ig g" using contour_integral_unique by simp
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2313
  also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1))
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2314
      + ?ig (reversepath l)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2315
    unfolding g_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2316
    by (auto simp add:contour_integrable_reversepath_eq)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2317
  also have "... = ?ig (circlepath z e2)  - ?ig (circlepath z e1)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2318
    by (auto simp add:contour_integral_reversepath)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2319
  finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2320
    by simp
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2321
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2322
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2323
lemma base_residue:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2324
  assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2325
    and r_cball:"cball z r \<subseteq> s"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  2326
  shows "(f has_contour_integral 2 * pi * \<i> * (residue f z)) (circlepath z r)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2327
proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2328
  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2329
    using open_contains_cball[of s] \<open>open s\<close> \<open>z\<in>s\<close> by auto
63589
58aab4745e85 more symbols;
wenzelm
parents: 63540
diff changeset
  2330
  define c where "c \<equiv> 2 * pi * \<i>"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2331
  define i where "i \<equiv> contour_integral (circlepath z e) f / c"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2332
  have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2333
    proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2334
      have "contour_integral (circlepath z e) f = contour_integral (circlepath z \<epsilon>) f"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2335
          "f contour_integrable_on circlepath z \<epsilon>"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2336
          "f contour_integrable_on circlepath z e"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2337
        using \<open>\<epsilon><e\<close>
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2338
        by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2339
      then show ?thesis unfolding i_def c_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2340
        by (auto intro:has_contour_integral_integral)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2341
    qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2342
  then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2343
    unfolding residue_def c_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2344
    apply (rule_tac someI[of _ i],intro  exI[where x=e])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2345
    by (auto simp add:\<open>e>0\<close> c_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2346
  then obtain e' where "e'>0"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2347
      and e'_def:"\<forall>\<epsilon>>0. \<epsilon><e' \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2348
    by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2349
  let ?int="\<lambda>e. contour_integral (circlepath z e) f"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2350
  define  \<epsilon> where "\<epsilon> \<equiv> Min {r,e'} / 2"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2351
  have "\<epsilon>>0" "\<epsilon>\<le>r" "\<epsilon><e'" using \<open>r>0\<close> \<open>e'>0\<close> unfolding \<epsilon>_def by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2352
  have "(f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2353
    using e'_def[rule_format,OF \<open>\<epsilon>>0\<close> \<open>\<epsilon><e'\<close>] .
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2354
  then show ?thesis unfolding c_def
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2355
    using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball]
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2356
    by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \<epsilon>" "circlepath z r"])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2357
qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2358
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2359
lemma residue_holo:
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2360
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2361
  shows "residue f z = 0"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2362
proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2363
  define c where "c \<equiv> 2 * pi * \<i>"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2364
  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2365
    using open_contains_cball_eq by blast
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2366
  have "(f has_contour_integral c*residue f z) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2367
    using f_holo
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2368
    by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2369
  moreover have "(f has_contour_integral 0) (circlepath z e)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2370
    using f_holo e_cball \<open>e>0\<close>
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2371
    by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2372
  ultimately have "c*residue f z =0"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2373
    using has_contour_integral_unique by blast
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2374
  thus ?thesis unfolding c_def  by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2375
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2376
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2377
lemma residue_const:"residue (\<lambda>_. c) z = 0"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2378
  by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2379
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2380
lemma residue_add:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2381
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2382
      and g_holo:"g holomorphic_on s - {z}"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2383
  shows "residue (\<lambda>z. f z + g z) z= residue f z + residue g z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2384
proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2385
  define c where "c \<equiv> 2 * pi * \<i>"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2386
  define fg where "fg \<equiv> (\<lambda>z. f z+g z)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2387
  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2388
    using open_contains_cball_eq by blast
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2389
  have "(fg has_contour_integral c * residue fg z) (circlepath z e)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2390
    unfolding fg_def using f_holo g_holo
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2391
    apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2392
    by (auto intro:holomorphic_intros)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2393
  moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2394
    unfolding fg_def using f_holo g_holo
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2395
    by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2396
  ultimately have "c*(residue f z + residue g z) = c * residue fg z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2397
    using has_contour_integral_unique by (auto simp add:distrib_left)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2398
  thus ?thesis unfolding fg_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2399
    by (auto simp add:c_def)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2400
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2401
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2402
lemma residue_lmul:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2403
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2404
  shows "residue (\<lambda>z. c * (f z)) z= c * residue f z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2405
proof (cases "c=0")
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2406
  case True
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2407
  thus ?thesis using residue_const by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2408
next
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2409
  case False
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2410
  define c' where "c' \<equiv> 2 * pi * \<i>"
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2411
  define f' where "f' \<equiv> (\<lambda>z. c * (f z))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2412
  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2413
    using open_contains_cball_eq by blast
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2414
  have "(f' has_contour_integral c' * residue f' z) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2415
    unfolding f'_def using f_holo
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2416
    apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2417
    by (auto intro:holomorphic_intros)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2418
  moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2419
    unfolding f'_def using f_holo
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2420
    by (auto intro: has_contour_integral_lmul
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2421
      base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2422
  ultimately have "c' * residue f' z  = c * (c' * residue f z)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2423
    using has_contour_integral_unique by auto
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2424
  thus ?thesis unfolding f'_def c'_def using False
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2425
    by (auto simp add:field_simps)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2426
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2427
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2428
lemma residue_rmul:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2429
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2430
  shows "residue (\<lambda>z. (f z) * c) z= residue f z * c"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2431
using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2432
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2433
lemma residue_div:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2434
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2435
  shows "residue (\<lambda>z. (f z) / c) z= residue f z / c "
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2436
using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2437
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2438
lemma residue_neg:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2439
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2440
  shows "residue (\<lambda>z. - (f z)) z= - residue f z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2441
using residue_lmul[OF assms,of "-1"] by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2442
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2443
lemma residue_diff:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2444
  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2445
      and g_holo:"g holomorphic_on s - {z}"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2446
  shows "residue (\<lambda>z. f z - g z) z= residue f z - residue g z"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2447
using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)]
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2448
by (auto intro:holomorphic_intros g_holo)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2449
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2450
lemma residue_simple:
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2451
  assumes "open s" "z\<in>s" and f_holo:"f holomorphic_on s"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2452
  shows "residue (\<lambda>w. f w / (w - z)) z = f z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2453
proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2454
  define c where "c \<equiv> 2 * pi * \<i>"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2455
  define f' where "f' \<equiv> \<lambda>w. f w / (w - z)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2456
  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2457
    using open_contains_cball_eq by blast
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2458
  have "(f' has_contour_integral c * f z) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2459
    unfolding f'_def c_def using \<open>e>0\<close> f_holo e_cball
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2460
    by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2461
  moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2462
    unfolding f'_def using f_holo
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2463
    apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2464
    by (auto intro!:holomorphic_intros)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2465
  ultimately have "c * f z = c * residue f' z"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2466
    using has_contour_integral_unique by blast
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2467
  thus ?thesis unfolding c_def f'_def  by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2468
qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2469
66286
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2470
lemma residue_simple':
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2471
  assumes s: "open s" "z \<in> s" and holo: "f holomorphic_on (s - {z})" 
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2472
      and lim: "((\<lambda>w. f w * (w - z)) \<longlongrightarrow> c) (at z)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2473
  shows   "residue f z = c"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2474
proof -
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2475
  define g where "g = (\<lambda>w. if w = z then c else f w * (w - z))"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2476
  from holo have "(\<lambda>w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P")
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2477
    by (force intro: holomorphic_intros)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2478
  also have "?P \<longleftrightarrow> g holomorphic_on (s - {z})"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2479
    by (intro holomorphic_cong refl) (simp_all add: g_def)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2480
  finally have *: "g holomorphic_on (s - {z})" .
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2481
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2482
  note lim
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2483
  also have "(\<lambda>w. f w * (w - z)) \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z\<rightarrow> g z"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2484
    by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2485
  finally have **: "g \<midarrow>z\<rightarrow> g z" .
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2486
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2487
  have g_holo: "g holomorphic_on s"
68359
8cd3d0305269 tidied more Cauchy proofs
paulson <lp15@cam.ac.uk>
parents: 68310
diff changeset
  2488
    by (rule no_isolated_singularity'[where K = "{z}"])
66286
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2489
       (insert assms * **, simp_all add: at_within_open_NO_MATCH)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2490
  from s and this have "residue (\<lambda>w. g w / (w - z)) z = g z"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2491
    by (rule residue_simple)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2492
  also have "\<forall>\<^sub>F za in at z. g za / (za - z) = f za"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2493
    unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2494
  hence "residue (\<lambda>w. g w / (w - z)) z = residue f z"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2495
    by (intro residue_cong refl)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2496
  finally show ?thesis
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2497
    by (simp add: g_def)
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2498
qed
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65965
diff changeset
  2499
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2500
lemma residue_holomorphic_over_power:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2501
  assumes "open A" "z0 \<in> A" "f holomorphic_on A"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2502
  shows   "residue (\<lambda>z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2503
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2504
  let ?f = "\<lambda>z. f z / (z - z0) ^ Suc n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2505
  from assms(1,2) obtain r where r: "r > 0" "cball z0 r \<subseteq> A"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2506
    by (auto simp: open_contains_cball)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2507
  have "(?f has_contour_integral 2 * pi * \<i> * residue ?f z0) (circlepath z0 r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2508
    using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2509
  moreover have "(?f has_contour_integral 2 * pi * \<i> / fact n * (deriv ^^ n) f z0) (circlepath z0 r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2510
    using assms r
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2511
    by (intro Cauchy_has_contour_integral_higher_derivative_circlepath)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2512
       (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2513
  ultimately have "2 * pi * \<i> * residue ?f z0 = 2 * pi * \<i> / fact n * (deriv ^^ n) f z0"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2514
    by (rule has_contour_integral_unique)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2515
  thus ?thesis by (simp add: field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2516
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2517
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2518
lemma residue_holomorphic_over_power':
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2519
  assumes "open A" "0 \<in> A" "f holomorphic_on A"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2520
  shows   "residue (\<lambda>z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2521
  using residue_holomorphic_over_power[OF assms] by simp
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2522
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2523
lemma get_integrable_path:
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2524
  assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2525
  obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2526
    "path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2527
proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>])
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2528
  case 1
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2529
  obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2530
    using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2531
      valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2532
  moreover have "f contour_integrable_on g"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2533
    using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f]
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2534
      \<open>f holomorphic_on s - {}\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2535
    by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2536
  ultimately show ?case using "1"(1)[of g] by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2537
next
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2538
  case idt:(2 p pts)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2539
  obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2540
    using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a]
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2541
      \<open>a \<in> s - insert p pts\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2542
    by auto
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2543
  define a' where "a' \<equiv> a+e/2"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2544
  have "a'\<in>s-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2545
    by (auto simp add:dist_complex_def a'_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2546
  then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2547
    "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2548
    using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2549
    by (metis Diff_insert2 open_delete)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2550
  define g where "g \<equiv> linepath a a' +++ g'"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2551
  have "valid_path g" unfolding g_def by (auto intro: valid_path_join)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2552
  moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2553
  moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2554
    proof (rule subset_path_image_join)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2555
      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2556
        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2557
      then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2558
        by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2559
    next
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2560
      show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2561
    qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2562
  moreover have "f contour_integrable_on g"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2563
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2564
      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2565
        by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2566
      then have "continuous_on (closed_segment a a') f"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2567
        using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)]
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2568
        apply (elim continuous_on_subset)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2569
        by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2570
      then have "f contour_integrable_on linepath a a'"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2571
        using contour_integrable_continuous_linepath by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2572
      then show ?thesis unfolding g_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2573
        apply (rule contour_integrable_joinI)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2574
        by (auto simp add: \<open>e>0\<close>)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2575
    qed
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2576
  ultimately show ?case using idt.prems(1)[of g] by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2577
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2578
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2579
lemma Cauchy_theorem_aux:
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2580
  assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2581
          "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2582
          "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2583
          "\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2584
  shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2585
    using assms
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2586
proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>])
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2587
  case 1
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2588
  then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2589
next
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2590
  case (2 p pts)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2591
  note fin[simp] = \<open>finite (insert p pts)\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2592
    and connected = \<open>connected (s - insert p pts)\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2593
    and valid[simp] = \<open>valid_path g\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2594
    and g_loop[simp] = \<open>pathfinish g = pathstart g\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2595
    and holo[simp]= \<open>f holomorphic_on s - insert p pts\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2596
    and path_img = \<open>path_image g \<subseteq> s - insert p pts\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2597
    and winding = \<open>\<forall>z. z \<notin> s \<longrightarrow> winding_number g z = 0\<close>
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2598
    and h = \<open>\<forall>pa\<in>s. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s \<and> (w \<noteq> pa \<longrightarrow> w \<notin> insert p pts))\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2599
  have "h p>0" and "p\<in>s"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2600
    and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2601
    using h \<open>insert p pts \<subseteq> s\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2602
  obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2603
      "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2604
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2605
      have "p + h p\<in>cball p (h p)" using h[rule_format,of p]
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2606
        by (simp add: \<open>p \<in> s\<close> dist_norm)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2607
      then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2608
        by fastforce
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2609
      moreover have "pathstart g \<in> s - insert p pts " using path_img by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2610
      ultimately show ?thesis
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2611
        using get_integrable_path[OF \<open>open s\<close> connected fin holo,of "pathstart g" "p+h p"] that
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2612
        by blast
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2613
    qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2614
  obtain n::int where "n=winding_number g p"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2615
    using integer_winding_number[OF _ g_loop,of p] valid path_img
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2616
    by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2617
  define p_circ where "p_circ \<equiv> circlepath p (h p)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2618
  define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67135
diff changeset
  2619
  define n_circ where "n_circ \<equiv> \<lambda>n. ((+++) p_circ ^^ n) p_circ_pt"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2620
  define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2621
  have n_circ:"valid_path (n_circ k)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2622
      "winding_number (n_circ k) p = k"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2623
      "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2624
      "path_image (n_circ k) =  (if k=0 then {p + h p} else sphere p (h p))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2625
      "p \<notin> path_image (n_circ k)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2626
      "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2627
      "f contour_integrable_on (n_circ k)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2628
      "contour_integral (n_circ k) f = k *  contour_integral p_circ f"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2629
      for k
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2630
    proof (induct k)
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2631
      case 0
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2632
      show "valid_path (n_circ 0)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2633
        and "path_image (n_circ 0) =  (if 0=0 then {p + h p} else sphere p (h p))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2634
        and "winding_number (n_circ 0) p = of_nat 0"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2635
        and "pathstart (n_circ 0) = p + h p"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2636
        and "pathfinish (n_circ 0) = p + h p"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2637
        and "p \<notin> path_image (n_circ 0)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2638
        unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2639
        by (auto simp add: dist_norm)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2640
      show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p'
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2641
        unfolding n_circ_def p_circ_pt_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2642
        apply (auto intro!:winding_number_trivial)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2643
        by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2644
      show "f contour_integrable_on (n_circ 0)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2645
        unfolding n_circ_def p_circ_pt_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2646
        by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2647
      show "contour_integral (n_circ 0) f = of_nat 0  *  contour_integral p_circ f"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2648
        unfolding n_circ_def p_circ_pt_def by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2649
    next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2650
      case (Suc k)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2651
      have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2652
      have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2653
        using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2654
      have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2655
        proof -
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2656
          have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2657
          then show ?thesis using h_p pcirc(1) by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2658
        qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2659
      have pcirc_integrable:"f contour_integrable_on p_circ"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2660
        by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def]
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2661
          contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2662
          holomorphic_on_subset[OF holo])
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2663
      show "valid_path (n_circ (Suc k))"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2664
        using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2665
      show "path_image (n_circ (Suc k))
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2666
          = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2667
        proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2668
          have "path_image p_circ = sphere p (h p)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2669
            unfolding p_circ_def using \<open>0 < h p\<close> by auto
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2670
          then show ?thesis unfolding n_Suc  using Suc.hyps(5)  \<open>h p>0\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2671
            by (auto simp add:  path_image_join[OF pcirc(3)]  dist_norm)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2672
        qed
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2673
      then show "p \<notin> path_image (n_circ (Suc k))" using \<open>h p>0\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2674
      show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2675
        proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2676
          have "winding_number p_circ p = 1"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2677
            by (simp add: \<open>h p > 0\<close> p_circ_def winding_number_circlepath_centre)
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2678
          moreover have "p \<notin> path_image (n_circ k)" using Suc(5) \<open>h p>0\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2679
          then have "winding_number (p_circ +++ n_circ k) p
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2680
              = winding_number p_circ p + winding_number (n_circ k) p"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2681
            using  valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2682
            apply (intro winding_number_join)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2683
            by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2684
          ultimately show ?thesis using Suc(2) unfolding n_circ_def
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2685
            by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2686
        qed
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2687
      show "pathstart (n_circ (Suc k)) = p + h p"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2688
        by (simp add: n_circ_def p_circ_def)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2689
      show "pathfinish (n_circ (Suc k)) = p + h p"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2690
        using Suc(4) unfolding n_circ_def by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2691
      show "winding_number (n_circ (Suc k)) p'=0 \<and>  p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p'
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2692
        proof -
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2693
          have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2694
          moreover have "p' \<notin> path_image (n_circ k)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2695
            using Suc.hyps(7) that by blast
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2696
          moreover have "winding_number p_circ p' = 0"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2697
            proof -
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2698
              have "path_image p_circ \<subseteq> cball p (h p)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2699
                using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2700
              moreover have "p'\<notin>cball p (h p)" using \<open>p \<in> s\<close> h that "2.hyps"(2) by fastforce
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2701
              ultimately show ?thesis unfolding p_circ_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2702
                apply (intro winding_number_zero_outside)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2703
                by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2704
            qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2705
          ultimately show ?thesis
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2706
            unfolding n_Suc
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2707
            apply (subst winding_number_join)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2708
            by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that])
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2709
        qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2710
      show "f contour_integrable_on (n_circ (Suc k))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2711
        unfolding n_Suc
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2712
        by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)])
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2713
      show "contour_integral (n_circ (Suc k)) f = (Suc k) *  contour_integral p_circ f"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2714
        unfolding n_Suc
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2715
        by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2716
          Suc(9) algebra_simps)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2717
    qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2718
  have cp[simp]:"pathstart cp = p + h p"  "pathfinish cp = p + h p"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2719
         "valid_path cp" "path_image cp \<subseteq> s - insert p pts"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2720
         "winding_number cp p = - n"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2721
         "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2722
         "f contour_integrable_on cp"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2723
         "contour_integral cp f = - n * contour_integral p_circ f"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2724
    proof -
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2725
      show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2726
        using n_circ unfolding cp_def by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2727
    next
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2728
      have "sphere p (h p) \<subseteq>  s - insert p pts"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2729
        using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2730
      moreover  have "p + complex_of_real (h p) \<in> s - insert p pts"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2731
        using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2732
      ultimately show "path_image cp \<subseteq>  s - insert p pts" unfolding cp_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2733
        using n_circ(5)  by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2734
    next
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2735
      show "winding_number cp p = - n"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2736
        unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2737
        by (auto simp: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2738
    next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2739
      show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p'
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2740
        unfolding cp_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2741
        apply (auto)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2742
        apply (subst winding_number_reversepath)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2743
        by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1))
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2744
    next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2745
      show "f contour_integrable_on cp" unfolding cp_def
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2746
        using contour_integrable_reversepath_eq n_circ(1,8) by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2747
    next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2748
      show "contour_integral cp f = - n * contour_integral p_circ f"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2749
        unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2750
        by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2751
    qed
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2752
  define g' where "g' \<equiv> g +++ pg +++ cp +++ (reversepath pg)"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2753
  have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2754
    proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2755
      show "connected (s - {p} - pts)" using connected by (metis Diff_insert2)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2756
      show "open (s - {p})" using \<open>open s\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2757
      show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close>  by blast
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2758
      show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2759
      show "valid_path g'"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2760
        unfolding g'_def cp_def using n_circ valid pg g_loop
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2761
        by (auto intro!:valid_path_join )
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2762
      show "pathfinish g' = pathstart g'"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2763
        unfolding g'_def cp_def using pg(2) by simp
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2764
      show "path_image g' \<subseteq> s - {p} - pts"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2765
        proof -
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65039
diff changeset
  2766
          define s' where "s' \<equiv> s - {p} - pts"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2767
          have s':"s' = s-insert p pts " unfolding s'_def by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2768
          then show ?thesis using path_img pg(4) cp(4)
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2769
            unfolding g'_def
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2770
            apply (fold s'_def s')
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2771
            apply (intro subset_path_image_join)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2772
            by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2773
        qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2774
      note path_join_imp[simp]
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2775
      show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2776
        proof clarify
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2777
          fix z assume z:"z\<notin>s - {p}"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2778
          have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2779
              + winding_number (pg +++ cp +++ (reversepath pg)) z"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2780
            proof (rule winding_number_join)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2781
              show "path g" using \<open>valid_path g\<close> by (simp add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2782
              show "z \<notin> path_image g" using z path_img by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2783
              show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2784
                by (simp add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2785
            next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2786
              have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2787
                using pg(4) cp(4) by (auto simp:subset_path_image_join)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2788
              then show "z \<notin> path_image (pg +++ cp +++ reversepath pg)" using z by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2789
            next
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2790
              show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2791
            qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2792
          also have "... = winding_number g z + (winding_number pg z
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2793
              + winding_number (cp +++ (reversepath pg)) z)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2794
            proof (subst add_left_cancel,rule winding_number_join)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2795
              show "path pg" and "path (cp +++ reversepath pg)"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2796
               and "pathfinish pg = pathstart (cp +++ reversepath pg)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2797
                by (auto simp add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2798
              show "z \<notin> path_image pg" using pg(4) z by blast
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2799
              show "z \<notin> path_image (cp +++ reversepath pg)" using z
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2800
                by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2801
                  not_in_path_image_join path_image_reversepath singletonD)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2802
            qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2803
          also have "... = winding_number g z + (winding_number pg z
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2804
              + (winding_number cp z + winding_number (reversepath pg) z))"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2805
            apply (auto intro!:winding_number_join simp: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2806
            apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2807
            by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2808
          also have "... = winding_number g z + winding_number cp z"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2809
            apply (subst winding_number_reversepath)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2810
            apply (auto simp: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2811
            by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2812
          finally have "winding_number g' z = winding_number g z + winding_number cp z"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2813
            unfolding g'_def .
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2814
          moreover have "winding_number g z + winding_number cp z = 0"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2815
            using winding z \<open>n=winding_number g p\<close> by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2816
          ultimately show "winding_number g' z = 0" unfolding g'_def by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2817
        qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2818
      show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2819
        using h by fastforce
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2820
    qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2821
  moreover have "contour_integral g' f = contour_integral g f
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2822
      - winding_number g p * contour_integral p_circ f"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2823
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2824
      have "contour_integral g' f =  contour_integral g f
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2825
        + contour_integral (pg +++ cp +++ reversepath pg) f"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2826
        unfolding g'_def
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2827
        apply (subst contour_integral_join)
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2828
        by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]]
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2829
          intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2830
          contour_integrable_reversepath)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2831
      also have "... = contour_integral g f + contour_integral pg f
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2832
          + contour_integral (cp +++ reversepath pg) f"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2833
        apply (subst contour_integral_join)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2834
        by (auto simp add:contour_integrable_reversepath)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2835
      also have "... = contour_integral g f + contour_integral pg f
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2836
          + contour_integral cp f + contour_integral (reversepath pg) f"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2837
        apply (subst contour_integral_join)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2838
        by (auto simp add:contour_integrable_reversepath)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2839
      also have "... = contour_integral g f + contour_integral cp f"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2840
        using contour_integral_reversepath
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2841
        by (auto simp add:contour_integrable_reversepath)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2842
      also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f"
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2843
        using \<open>n=winding_number g p\<close> by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2844
      finally show ?thesis .
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2845
    qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2846
  moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p'
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2847
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2848
      have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2849
        using "2.prems"(8) that
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2850
        apply blast
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2851
        apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that)
69712
dc85b5b3a532 renamings and new material
paulson <lp15@cam.ac.uk>
parents: 69661
diff changeset
  2852
        by (meson DiffD2 cp(4) rev_subsetD subset_insertI that)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2853
      have "winding_number g' p' = winding_number g p'
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2854
          + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2855
        apply (subst winding_number_join)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2856
        apply (simp_all add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2857
        apply (intro not_in_path_image_join)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2858
        by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2859
      also have "... = winding_number g p' + winding_number pg p'
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2860
          + winding_number (cp +++ reversepath pg) p'"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2861
        apply (subst winding_number_join)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2862
        apply (simp_all add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2863
        apply (intro not_in_path_image_join)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2864
        by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2865
      also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p'
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2866
          + winding_number (reversepath pg) p'"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2867
        apply (subst winding_number_join)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2868
        by (simp_all add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2869
      also have "... = winding_number g p' + winding_number cp p'"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2870
        apply (subst winding_number_reversepath)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2871
        by (simp_all add: valid_path_imp_path)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2872
      also have "... = winding_number g p'" using that by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2873
      finally show ?thesis .
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2874
    qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2875
  ultimately show ?case unfolding p_circ_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2876
    apply (subst (asm) sum.cong[OF refl,
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2877
        of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2878
    by (auto simp add:sum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2879
qed
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2880
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2881
lemma Cauchy_theorem_singularities:
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2882
  assumes "open s" "connected s" "finite pts" and
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2883
          holo:"f holomorphic_on s-pts" and
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2884
          "valid_path g" and
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2885
          loop:"pathfinish g = pathstart g" and
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2886
          "path_image g \<subseteq> s-pts" and
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2887
          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" and
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2888
          avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2889
  shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2890
    (is "?L=?R")
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2891
proof -
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2892
  define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2893
  define pts1 where "pts1 \<equiv> pts \<inter> s"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2894
  define pts2 where "pts2 \<equiv> pts - pts1"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2895
  have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2896
    unfolding pts1_def pts2_def by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2897
  have "contour_integral g f =  (\<Sum>p\<in>pts1. circ p)" unfolding circ_def
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2898
    proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo])
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2899
      have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2900
      then show "connected (s - pts1)"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2901
        using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2902
    next
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2903
      show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2904
      show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2905
      show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2906
      show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2907
        by (simp add: avoid pts1_def)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2908
    qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2909
  moreover have "sum circ pts2=0"
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2910
    proof -
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2911
      have "winding_number g p=0" when "p\<in>pts2" for p
62837
237ef2bab6c7 isabelle update_cartouches -c -t;
wenzelm
parents: 62540
diff changeset
  2912
        using  \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2913
      thus ?thesis unfolding circ_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2914
        apply (intro sum.neutral)
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2915
        by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2916
    qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2917
  moreover have "?R=sum circ pts1 + sum circ pts2"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2918
    unfolding circ_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2919
    using sum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2920
    by blast
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2921
  ultimately show ?thesis
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2922
    apply (fold circ_def)
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2923
    by auto
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2924
qed
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  2925
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  2926
theorem Residue_theorem:
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2927
  fixes s pts::"complex set" and f::"complex \<Rightarrow> complex"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2928
    and g::"real \<Rightarrow> complex"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2929
  assumes "open s" "connected s" "finite pts" and
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2930
          holo:"f holomorphic_on s-pts" and
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2931
          "valid_path g" and
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2932
          loop:"pathfinish g = pathstart g" and
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2933
          "path_image g \<subseteq> s-pts" and
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2934
          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2935
  shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2936
proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2937
  define c where "c \<equiv>  2 * pi * \<i>"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2938
  obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2939
    using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2940
  have "contour_integral g f
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2941
      = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2942
    using Cauchy_theorem_singularities[OF assms avoid] .
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2943
  also have "... = (\<Sum>p\<in>pts.  c * winding_number g p * residue f p)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2944
    proof (intro sum.cong)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2945
      show "pts = pts" by simp
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2946
    next
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2947
      fix x assume "x \<in> pts"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2948
      show "winding_number g x * contour_integral (circlepath x (h x)) f
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2949
          = c * winding_number g x * residue f x"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2950
        proof (cases "x\<in>s")
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2951
          case False
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2952
          then have "winding_number g x=0" using homo by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2953
          thus ?thesis by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2954
        next
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2955
          case True
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2956
          have "contour_integral (circlepath x (h x)) f = c* residue f x"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2957
            using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True]
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2958
            apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2959
            by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2960
          then show ?thesis by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2961
        qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2962
    qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2963
  also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63928
diff changeset
  2964
    by (simp add: sum_distrib_left algebra_simps)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2965
  finally show ?thesis unfolding c_def .
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2966
qed
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2967
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2968
subsection \<open>Non-essential singular points\<close>
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2969
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
  2970
definition\<^marker>\<open>tag important\<close> is_pole :: 
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  2971
  "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2972
  "is_pole f a =  (LIM x (at a). f x :> at_infinity)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2973
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2974
lemma is_pole_cong:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2975
  assumes "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2976
  shows "is_pole f a \<longleftrightarrow> is_pole g b"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2977
  unfolding is_pole_def using assms by (intro filterlim_cong,auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2978
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2979
lemma is_pole_transform:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2980
  assumes "is_pole f a" "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2981
  shows "is_pole g b"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2982
  using is_pole_cong assms by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  2983
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2984
lemma is_pole_tendsto:
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2985
  fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2986
  shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2987
unfolding is_pole_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2988
by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2989
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2990
lemma is_pole_inverse_holomorphic:
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2991
  assumes "open s"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2992
    and f_holo:"f holomorphic_on (s-{z})"
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  2993
    and pole:"is_pole f z"
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2994
    and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2995
  shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2996
proof -
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2997
  define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2998
  have "isCont g z" unfolding isCont_def  using is_pole_tendsto[OF pole]
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  2999
    apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \<circ> f"])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3000
    by (simp_all add:g_def)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3001
  moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3002
  hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3003
    by (auto elim!:continuous_on_inverse simp add:non_z)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  3004
  hence "continuous_on (s-{z}) g" unfolding g_def
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3005
    apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3006
    by auto
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3007
  ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close>
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3008
    by (auto simp add:continuous_on_eq_continuous_at)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3009
  moreover have "(inverse o f) holomorphic_on (s-{z})"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3010
    unfolding comp_def using f_holo
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3011
    by (auto elim!:holomorphic_on_inverse simp add:non_z)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3012
  hence "g holomorphic_on (s-{z})"
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3013
    apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"])
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3014
    by (auto simp add:g_def)
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3015
  ultimately show ?thesis unfolding g_def using \<open>open s\<close>
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  3016
    by (auto elim!: no_isolated_singularity)
63151
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3017
qed
82df5181d699 updated proof of Residue Theorem (form Wenda Li)
paulson <lp15@cam.ac.uk>
parents: 63040
diff changeset
  3018
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3019
lemma not_is_pole_holomorphic:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3020
  assumes "open A" "x \<in> A" "f holomorphic_on A"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3021
  shows   "\<not>is_pole f x"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3022
proof -
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3023
  have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3024
  with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3025
  hence "f \<midarrow>x\<rightarrow> f x" by (simp add: isCont_def)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3026
  thus "\<not>is_pole f x" unfolding is_pole_def
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3027
    using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3028
qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66884
diff changeset
  3029
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3030
lemma is_pole_inverse_power: "n > 0 \<Longrightarrow> is_pole (\<lambda>z::complex. 1 / (z - a) ^ n) a"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3031
  unfolding is_pole_def inverse_eq_divide [symmetric]
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3032
  by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3033
     (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3034
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3035
lemma is_pole_inverse: "is_pole (\<lambda>z::complex. 1 / (z - a)) a"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3036
  using is_pole_inverse_power[of 1 a] by simp
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3037
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3038
lemma is_pole_divide:
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3039
  fixes f :: "'a :: t2_space \<Rightarrow> 'b :: real_normed_field"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3040
  assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \<noteq> 0"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3041
  shows   "is_pole (\<lambda>z. f z / g z) z"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3042
proof -
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3043
  have "filterlim (\<lambda>z. f z * inverse (g z)) at_infinity (at z)"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3044
    by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"]
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3045
                 filterlim_compose[OF filterlim_inverse_at_infinity])+
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3046
       (insert assms, auto simp: isCont_def)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3047
  thus ?thesis by (simp add: divide_simps is_pole_def)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3048
qed
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3049
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3050
lemma is_pole_basic:
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3051
  assumes "f holomorphic_on A" "open A" "z \<in> A" "f z \<noteq> 0" "n > 0"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3052
  shows   "is_pole (\<lambda>w. f w / (w - z) ^ n) z"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3053
proof (rule is_pole_divide)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3054
  have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3055
  with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3056
  have "filterlim (\<lambda>w. (w - z) ^ n) (nhds 0) (at z)"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3057
    using assms by (auto intro!: tendsto_eq_intros)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3058
  thus "filterlim (\<lambda>w. (w - z) ^ n) (at 0) (at z)"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3059
    by (intro filterlim_atI tendsto_eq_intros)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3060
       (insert assms, auto simp: eventually_at_filter)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3061
qed fact+
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3062
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3063
lemma is_pole_basic':
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3064
  assumes "f holomorphic_on A" "open A" "0 \<in> A" "f 0 \<noteq> 0" "n > 0"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3065
  shows   "is_pole (\<lambda>w. f w / w ^ n) 0"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3066
  using is_pole_basic[of f A 0] assms by simp
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3067
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3068
text \<open>The proposition 
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
  3069
              \<^term>\<open>\<exists>x. ((f::complex\<Rightarrow>complex) \<longlongrightarrow> x) (at z) \<or> is_pole f z\<close> 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
  3070
can be interpreted as the complex function \<^term>\<open>f\<close> has a non-essential singularity at \<^term>\<open>z\<close> 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3071
(i.e. the singularity is either removable or a pole).\<close> 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3072
definition not_essential::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3073
  "not_essential f z = (\<exists>x. f\<midarrow>z\<rightarrow>x \<or> is_pole f z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3074
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3075
definition isolated_singularity_at::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3076
  "isolated_singularity_at f z = (\<exists>r>0. f analytic_on ball z r-{z})"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3077
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3078
named_theorems singularity_intros "introduction rules for singularities"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3079
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3080
lemma holomorphic_factor_unique:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3081
  fixes f::"complex \<Rightarrow> complex" and z::complex and r::real and m n::int
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3082
  assumes "r>0" "g z\<noteq>0" "h z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3083
    and asm:"\<forall>w\<in>ball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0 \<and> f w =  h w * (w - z) powr m \<and> h w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3084
    and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3085
  shows "n=m"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3086
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3087
  have [simp]:"at z within ball z r \<noteq> bot" using \<open>r>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3088
      by (auto simp add:at_within_ball_bot_iff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3089
  have False when "n>m"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3090
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3091
    have "(h \<longlongrightarrow> 0) (at z within ball z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3092
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (n - m) * g w"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3093
      have "\<forall>w\<in>ball z r-{z}. h w = (w-z)powr(n-m) * g w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3094
        using \<open>n>m\<close> asm \<open>r>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3095
        apply (auto simp add:field_simps powr_diff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3096
        by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3097
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3098
            \<Longrightarrow> (x' - z) powr (n - m) * g x' = h x'" for x' by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3099
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3100
      define F where "F \<equiv> at z within ball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3101
      define f' where "f' \<equiv> \<lambda>x. (x - z) powr (n-m)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3102
      have "f' z=0" using \<open>n>m\<close> unfolding f'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3103
      moreover have "continuous F f'" unfolding f'_def F_def continuous_def
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  3104
        apply (subst Lim_ident_at)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3105
        using \<open>n>m\<close> by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros)  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3106
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3107
        by (simp add: continuous_within)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3108
      moreover have "(g \<longlongrightarrow> g z) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3109
        using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3110
        unfolding F_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3111
      ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3112
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3113
    moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3114
      using holomorphic_on_imp_continuous_on[OF h_holo]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3115
      by (auto simp add:continuous_on_def \<open>r>0\<close>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3116
    ultimately have "h z=0" by (auto intro!: tendsto_unique)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3117
    thus False using \<open>h z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3118
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3119
  moreover have False when "m>n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3120
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3121
    have "(g \<longlongrightarrow> 0) (at z within ball z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3122
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (m - n) * h w"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3123
      have "\<forall>w\<in>ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \<open>m>n\<close> asm
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3124
        apply (auto simp add:field_simps powr_diff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3125
        by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3126
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3127
            \<Longrightarrow> (x' - z) powr (m - n) * h x' = g x'" for x' by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3128
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3129
      define F where "F \<equiv> at z within ball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3130
      define f' where "f' \<equiv>\<lambda>x. (x - z) powr (m-n)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3131
      have "f' z=0" using \<open>m>n\<close> unfolding f'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3132
      moreover have "continuous F f'" unfolding f'_def F_def continuous_def
70065
cc89a395b5a3 Free_Abelian_Groups finally working; fixed some duplicates; cleaned up some proofs
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  3133
        apply (subst Lim_ident_at)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3134
        using \<open>m>n\<close> by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3135
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3136
        by (simp add: continuous_within)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3137
      moreover have "(h \<longlongrightarrow> h z) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3138
        using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3139
        unfolding F_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3140
      ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3141
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3142
    moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3143
      using holomorphic_on_imp_continuous_on[OF g_holo]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3144
      by (auto simp add:continuous_on_def \<open>r>0\<close>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3145
    ultimately have "g z=0" by (auto intro!: tendsto_unique)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3146
    thus False using \<open>g z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3147
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3148
  ultimately show "n=m" by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3149
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3150
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3151
lemma holomorphic_factor_puncture:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3152
  assumes f_iso:"isolated_singularity_at f z"   
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
  3153
      and "not_essential f z" \<comment> \<open>\<^term>\<open>f\<close> has either a removable singularity or a pole at \<^term>\<open>z\<close>\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
  3154
      and non_zero:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0" \<comment> \<open>\<^term>\<open>f\<close> will not be constantly zero in a neighbour of \<^term>\<open>z\<close>\<close>
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3155
  shows "\<exists>!n::int. \<exists>g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3156
          \<and> (\<forall>w\<in>cball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3157
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3158
  define P where "P = (\<lambda>f n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3159
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n)  \<and> g w\<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3160
  have imp_unique:"\<exists>!n::int. \<exists>g r. P f n g r" when "\<exists>n g r. P f n g r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3161
  proof (rule ex_ex1I[OF that])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3162
    fix n1 n2 :: int
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3163
    assume g1_asm:"\<exists>g1 r1. P f n1 g1 r1" and g2_asm:"\<exists>g2 r2. P f n2 g2 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3164
    define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r-{z}. f w = g w * (w - z) powr (of_int n) \<and> g w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3165
    obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3166
        and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3167
    obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3168
        and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3169
    define r where "r \<equiv> min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3170
    have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3171
    moreover have "\<forall>w\<in>ball z r-{z}. f w = g1 w * (w-z) powr n1 \<and> g1 w\<noteq>0 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3172
        \<and> f w = g2 w * (w - z) powr n2  \<and> g2 w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3173
      using \<open>fac n1 g1 r1\<close> \<open>fac n2 g2 r2\<close>   unfolding fac_def r_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3174
      by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3175
    ultimately show "n1=n2" using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3176
      apply (elim holomorphic_factor_unique)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3177
      by (auto simp add:r_def) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3178
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3179
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3180
  have P_exist:"\<exists> n g r. P h n g r" when 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3181
      "\<exists>z'. (h \<longlongrightarrow> z') (at z)" "isolated_singularity_at h z"  "\<exists>\<^sub>Fw in (at z). h w\<noteq>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3182
    for h
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3183
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3184
    from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3185
      unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3186
    obtain z' where "(h \<longlongrightarrow> z') (at z)" using \<open>\<exists>z'. (h \<longlongrightarrow> z') (at z)\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3187
    define h' where "h'=(\<lambda>x. if x=z then z' else h x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3188
    have "h' holomorphic_on ball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3189
      apply (rule no_isolated_singularity'[of "{z}"]) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3190
      subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \<open>h \<midarrow>z\<rightarrow> z'\<close> empty_iff h'_def insert_iff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3191
      subgoal using \<open>h analytic_on ball z r - {z}\<close> analytic_imp_holomorphic h'_def holomorphic_transform 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3192
        by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3193
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3194
    have ?thesis when "z'=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3195
    proof - 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3196
      have "h' z=0" using that unfolding h'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3197
      moreover have "\<not> h' constant_on ball z r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3198
        using \<open>\<exists>\<^sub>Fw in (at z). h w\<noteq>0\<close> unfolding constant_on_def frequently_def eventually_at h'_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3199
        apply simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3200
        by (metis \<open>0 < r\<close> centre_in_ball dist_commute mem_ball that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3201
      moreover note \<open>h' holomorphic_on ball z r\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3202
      ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \<subseteq> ball z r" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3203
          g:"g holomorphic_on ball z r1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3204
          "\<And>w. w \<in> ball z r1 \<Longrightarrow> h' w = (w - z) ^ n * g w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3205
          "\<And>w. w \<in> ball z r1 \<Longrightarrow> g w \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3206
        using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified,
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3207
                OF \<open>h' holomorphic_on ball z r\<close> \<open>r>0\<close> \<open>h' z=0\<close> \<open>\<not> h' constant_on ball z r\<close>] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3208
        by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3209
      define rr where "rr=r1/2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3210
      have "P h' n g rr"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3211
        unfolding P_def rr_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3212
        using \<open>n>0\<close> \<open>r1>0\<close> g by (auto simp add:powr_nat)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3213
      then have "P h n g rr"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3214
        unfolding h'_def P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3215
      then show ?thesis unfolding P_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3216
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3217
    moreover have ?thesis when "z'\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3218
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3219
      have "h' z\<noteq>0" using that unfolding h'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3220
      obtain r1 where "r1>0" "cball z r1 \<subseteq> ball z r" "\<forall>x\<in>cball z r1. h' x\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3221
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3222
        have "isCont h' z" "h' z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3223
          by (auto simp add: Lim_cong_within \<open>h \<midarrow>z\<rightarrow> z'\<close> \<open>z'\<noteq>0\<close> continuous_at h'_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3224
        then obtain r2 where r2:"r2>0" "\<forall>x\<in>ball z r2. h' x\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3225
          using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3226
        define r1 where "r1=min r2 r / 2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3227
        have "0 < r1" "cball z r1 \<subseteq> ball z r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3228
          using \<open>r2>0\<close> \<open>r>0\<close> unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3229
        moreover have "\<forall>x\<in>cball z r1. h' x \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3230
          using r2 unfolding r1_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3231
        ultimately show ?thesis using that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3232
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3233
      then have "P h' 0 h' r1" using \<open>h' holomorphic_on ball z r\<close> unfolding P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3234
      then have "P h 0 h' r1" unfolding P_def h'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3235
      then show ?thesis unfolding P_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3236
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3237
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3238
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3239
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3240
  have ?thesis when "\<exists>x. (f \<longlongrightarrow> x) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3241
    apply (rule_tac imp_unique[unfolded P_def])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3242
    using P_exist[OF that(1) f_iso non_zero] unfolding P_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3243
  moreover have ?thesis when "is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3244
  proof (rule imp_unique[unfolded P_def])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3245
    obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3246
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3247
      have "\<forall>\<^sub>F z in at z. f z \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3248
        using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3249
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3250
      then obtain e1 where e1:"e1>0" "\<forall>x\<in>ball z e1-{z}. f x\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3251
        using that eventually_at[of "\<lambda>x. f x\<noteq>0" z UNIV,simplified] by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3252
      obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3253
        using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3254
      define e where "e=min e1 e2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3255
      show ?thesis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3256
        apply (rule that[of e])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3257
        using  e1 e2 unfolding e_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3258
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3259
    
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3260
    define h where "h \<equiv> \<lambda>x. inverse (f x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3261
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3262
    have "\<exists>n g r. P h n g r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3263
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3264
      have "h \<midarrow>z\<rightarrow> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3265
        using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3266
      moreover have "\<exists>\<^sub>Fw in (at z). h w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3267
        using non_zero 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3268
        apply (elim frequently_rev_mp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3269
        unfolding h_def eventually_at by (auto intro:exI[where x=1])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3270
      moreover have "isolated_singularity_at h z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3271
        unfolding isolated_singularity_at_def h_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3272
        apply (rule exI[where x=e])
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  3273
        using e_holo e_nz \<open>e>0\<close> by (metis open_ball analytic_on_open 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3274
            holomorphic_on_inverse open_delete)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3275
      ultimately show ?thesis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3276
        using P_exist[of h] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3277
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3278
    then obtain n g r
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3279
      where "0 < r" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3280
            g_holo:"g holomorphic_on cball z r" and "g z\<noteq>0" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3281
            g_fac:"(\<forall>w\<in>cball z r-{z}. h w = g w * (w - z) powr of_int n  \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3282
      unfolding P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3283
    have "P f (-n) (inverse o g) r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3284
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3285
      have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\<in>cball z r - {z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3286
        using g_fac[rule_format,of w] that unfolding h_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3287
        apply (auto simp add:powr_minus )
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3288
        by (metis inverse_inverse_eq inverse_mult_distrib)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3289
      then show ?thesis 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3290
        unfolding P_def comp_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3291
        using \<open>r>0\<close> g_holo g_fac \<open>g z\<noteq>0\<close> by (auto intro:holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3292
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3293
    then show "\<exists>x g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z \<noteq> 0 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3294
                  \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int x  \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3295
      unfolding P_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3296
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3297
  ultimately show ?thesis using \<open>not_essential f z\<close> unfolding not_essential_def  by presburger
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3298
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3299
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3300
lemma not_essential_transform:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3301
  assumes "not_essential g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3302
  assumes "\<forall>\<^sub>F w in (at z). g w = f w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3303
  shows "not_essential f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3304
  using assms unfolding not_essential_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3305
  by (simp add: filterlim_cong is_pole_cong)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3306
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3307
lemma isolated_singularity_at_transform:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3308
  assumes "isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3309
  assumes "\<forall>\<^sub>F w in (at z). g w = f w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3310
  shows "isolated_singularity_at f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3311
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3312
  obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3313
    using assms(1) unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3314
  obtain r2 where "r2>0" and r2:" \<forall>x. x \<noteq> z \<and> dist x z < r2 \<longrightarrow> g x = f x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3315
    using assms(2) unfolding eventually_at by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3316
  define r3 where "r3=min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3317
  have "r3>0" unfolding r3_def using \<open>r1>0\<close> \<open>r2>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3318
  moreover have "f analytic_on ball z r3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3319
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3320
    have "g holomorphic_on ball z r3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3321
      using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3322
    then have "f holomorphic_on ball z r3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3323
      using r2 unfolding r3_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3324
      by (auto simp add:dist_commute elim!:holomorphic_transform)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3325
    then show ?thesis by (subst analytic_on_open,auto)  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3326
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3327
  ultimately show ?thesis unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3328
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3329
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3330
lemma not_essential_powr[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3331
  assumes "LIM w (at z). f w :> (at x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3332
  shows "not_essential (\<lambda>w. (f w) powr (of_int n)) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3333
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3334
  define fp where "fp=(\<lambda>w. (f w) powr (of_int n))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3335
  have ?thesis when "n>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3336
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3337
    have "(\<lambda>w.  (f w) ^ (nat n)) \<midarrow>z\<rightarrow> x ^ nat n" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3338
      using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3339
    then have "fp \<midarrow>z\<rightarrow> x ^ nat n" unfolding fp_def      
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3340
      apply (elim Lim_transform_within[where d=1],simp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3341
      by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3342
    then show ?thesis unfolding not_essential_def fp_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3343
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3344
  moreover have ?thesis when "n=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3345
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3346
    have "fp \<midarrow>z\<rightarrow> 1 " 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3347
      apply (subst tendsto_cong[where g="\<lambda>_.1"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3348
      using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3349
    then show ?thesis unfolding fp_def not_essential_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3350
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3351
  moreover have ?thesis when "n<0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3352
  proof (cases "x=0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3353
    case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3354
    have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3355
      apply (subst filterlim_inverse_at_iff[symmetric],simp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3356
      apply (rule filterlim_atI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3357
      subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3358
      subgoal using filterlim_at_within_not_equal[OF assms,of 0] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3359
        by (eventually_elim,insert that,auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3360
      done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3361
    then have "LIM w (at z). fp w :> at_infinity"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3362
    proof (elim filterlim_mono_eventually)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3363
      show "\<forall>\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3364
        using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3365
        apply eventually_elim
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3366
        using powr_of_int that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3367
    qed auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3368
    then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3369
  next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3370
    case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3371
    let ?xx= "inverse (x ^ (nat (-n)))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3372
    have "(\<lambda>w. inverse ((f w) ^ (nat (-n)))) \<midarrow>z\<rightarrow>?xx"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3373
      using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3374
    then have "fp \<midarrow>z\<rightarrow>?xx"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3375
      apply (elim Lim_transform_within[where d=1],simp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3376
      unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3377
          not_le power_eq_0_iff powr_0 powr_of_int that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3378
    then show ?thesis unfolding fp_def not_essential_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3379
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3380
  ultimately show ?thesis by linarith
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3381
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3382
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3383
lemma isolated_singularity_at_powr[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3384
  assumes "isolated_singularity_at f z" "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3385
  shows "isolated_singularity_at (\<lambda>w. (f w) powr (of_int n)) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3386
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3387
  obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3388
    using assms(1) unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3389
  then have r1:"f holomorphic_on ball z r1 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3390
    using analytic_on_open[of "ball z r1-{z}" f] by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3391
  obtain r2 where "r2>0" and r2:"\<forall>w. w \<noteq> z \<and> dist w z < r2 \<longrightarrow> f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3392
    using assms(2) unfolding eventually_at by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3393
  define r3 where "r3=min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3394
  have "(\<lambda>w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3395
    apply (rule holomorphic_on_powr_of_int)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3396
    subgoal unfolding r3_def using r1 by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3397
    subgoal unfolding r3_def using r2 by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3398
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3399
  moreover have "r3>0" unfolding r3_def using \<open>0 < r1\<close> \<open>0 < r2\<close> by linarith
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3400
  ultimately show ?thesis unfolding isolated_singularity_at_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3401
    apply (subst (asm) analytic_on_open[symmetric])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3402
    by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3403
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3404
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3405
lemma non_zero_neighbour:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3406
  assumes f_iso:"isolated_singularity_at f z"   
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3407
      and f_ness:"not_essential f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3408
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3409
    shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3410
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3411
  obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3412
          and fr: "fp holomorphic_on cball z fr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3413
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3414
    using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3415
  have "f w \<noteq> 0" when " w \<noteq> z" "dist w z < fr" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3416
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3417
    have "f w = fp w * (w - z) powr of_int fn" "fp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3418
      using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3419
    moreover have "(w - z) powr of_int fn \<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3420
      unfolding powr_eq_0_iff using \<open>w\<noteq>z\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3421
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3422
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3423
  then show ?thesis using \<open>fr>0\<close> unfolding eventually_at by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3424
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3425
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3426
lemma non_zero_neighbour_pole:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3427
  assumes "is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3428
  shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3429
  using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0]  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3430
  unfolding is_pole_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3431
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3432
lemma non_zero_neighbour_alt:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3433
  assumes holo: "f holomorphic_on S"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3434
      and "open S" "connected S" "z \<in> S"  "\<beta> \<in> S" "f \<beta> \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3435
    shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0 \<and> w\<in>S"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3436
proof (cases "f z = 0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3437
  case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3438
  from isolated_zeros[OF holo \<open>open S\<close> \<open>connected S\<close> \<open>z \<in> S\<close> True \<open>\<beta> \<in> S\<close> \<open>f \<beta> \<noteq> 0\<close>] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3439
  obtain r where "0 < r" "ball z r \<subseteq> S" "\<forall>w \<in> ball z r - {z}.f w \<noteq> 0" by metis 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3440
  then show ?thesis unfolding eventually_at 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3441
    apply (rule_tac x=r in exI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3442
    by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3443
next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3444
  case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3445
  obtain r1 where r1:"r1>0" "\<forall>y. dist z y < r1 \<longrightarrow> f y \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3446
    using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3447
      holo holomorphic_on_imp_continuous_on by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3448
  obtain r2 where r2:"r2>0" "ball z r2 \<subseteq> S" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3449
    using assms(2) assms(4) openE by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3450
  show ?thesis unfolding eventually_at 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3451
    apply (rule_tac x="min r1 r2" in exI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3452
    using r1 r2 by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3453
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3454
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3455
lemma not_essential_times[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3456
  assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3457
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3458
  shows "not_essential (\<lambda>w. f w * g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3459
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3460
  define fg where "fg = (\<lambda>w. f w * g w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3461
  have ?thesis when "\<not> ((\<exists>\<^sub>Fw in (at z). f w\<noteq>0) \<and> (\<exists>\<^sub>Fw in (at z). g w\<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3462
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3463
    have "\<forall>\<^sub>Fw in (at z). fg w=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3464
      using that[unfolded frequently_def, simplified] unfolding fg_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3465
      by (auto elim: eventually_rev_mp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3466
    from tendsto_cong[OF this] have "fg \<midarrow>z\<rightarrow>0" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3467
    then show ?thesis unfolding not_essential_def fg_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3468
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3469
  moreover have ?thesis when f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0" and g_nconst:"\<exists>\<^sub>Fw in (at z). g w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3470
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3471
    obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3472
          and fr: "fp holomorphic_on cball z fr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3473
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3474
      using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3475
    obtain gn gp gr where [simp]:"gp z \<noteq> 0" and "gr > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3476
          and gr: "gp holomorphic_on cball z gr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3477
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3478
      using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3479
  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3480
    define r1 where "r1=(min fr gr)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3481
    have "r1>0" unfolding r1_def using  \<open>fr>0\<close> \<open>gr>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3482
    have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3483
      when "w\<in>ball z r1 - {z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3484
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3485
      have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3486
        using fr(2)[rule_format,of w] that unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3487
      moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3488
        using gr(2)[rule_format, of w] that unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3489
      ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3490
        unfolding fg_def by (auto simp add:powr_add)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3491
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3492
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3493
    have [intro]: "fp \<midarrow>z\<rightarrow>fp z" "gp \<midarrow>z\<rightarrow>gp z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3494
        using fr(1) \<open>fr>0\<close> gr(1) \<open>gr>0\<close>
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  3495
        by (meson open_ball ball_subset_cball centre_in_ball 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3496
            continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3497
            holomorphic_on_subset)+
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3498
    have ?thesis when "fn+gn>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3499
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3500
      have "(\<lambda>w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \<midarrow>z\<rightarrow>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3501
        using that by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3502
      then have "fg \<midarrow>z\<rightarrow> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3503
        apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3504
        by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3505
              eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3506
              that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3507
      then show ?thesis unfolding not_essential_def fg_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3508
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3509
    moreover have ?thesis when "fn+gn=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3510
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3511
      have "(\<lambda>w. fp w * gp w) \<midarrow>z\<rightarrow>fp z*gp z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3512
        using that by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3513
      then have "fg \<midarrow>z\<rightarrow> fp z*gp z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3514
        apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3515
        apply (subst fg_times)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3516
        by (auto simp add:dist_commute that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3517
      then show ?thesis unfolding not_essential_def fg_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3518
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3519
    moreover have ?thesis when "fn+gn<0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3520
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3521
      have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3522
        apply (rule filterlim_divide_at_infinity)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3523
        apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3524
        using eventually_at_topological by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3525
      then have "is_pole fg z" unfolding is_pole_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3526
        apply (elim filterlim_transform_within[OF _ _ \<open>r1>0\<close>],simp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3527
        apply (subst fg_times,simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3528
        apply (subst powr_of_int)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3529
        using that by (auto simp add:divide_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3530
      then show ?thesis unfolding not_essential_def fg_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3531
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3532
    ultimately show ?thesis unfolding not_essential_def fg_def by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3533
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3534
  ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3535
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3536
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3537
lemma not_essential_inverse[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3538
  assumes f_ness:"not_essential f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3539
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3540
  shows "not_essential (\<lambda>w. inverse (f w)) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3541
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3542
  define vf where "vf = (\<lambda>w. inverse (f w))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3543
  have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3544
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3545
    have "\<forall>\<^sub>Fw in (at z). f w=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3546
      using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3547
    then have "\<forall>\<^sub>Fw in (at z). vf w=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3548
      unfolding vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3549
    from tendsto_cong[OF this] have "vf \<midarrow>z\<rightarrow>0" unfolding vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3550
    then show ?thesis unfolding not_essential_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3551
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3552
  moreover have ?thesis when "is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3553
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3554
    have "vf \<midarrow>z\<rightarrow>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3555
      using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3556
    then show ?thesis unfolding not_essential_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3557
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3558
  moreover have ?thesis when "\<exists>x. f\<midarrow>z\<rightarrow>x " and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3559
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3560
    from that obtain fz where fz:"f\<midarrow>z\<rightarrow>fz" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3561
    have ?thesis when "fz=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3562
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3563
      have "(\<lambda>w. inverse (vf w)) \<midarrow>z\<rightarrow>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3564
        using fz that unfolding vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3565
      moreover have "\<forall>\<^sub>F w in at z. inverse (vf w) \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3566
        using non_zero_neighbour[OF f_iso f_ness f_nconst]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3567
        unfolding vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3568
      ultimately have "is_pole vf z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3569
        using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3570
      then show ?thesis unfolding not_essential_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3571
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3572
    moreover have ?thesis when "fz\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3573
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3574
      have "vf \<midarrow>z\<rightarrow>inverse fz"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3575
        using fz that unfolding vf_def by (auto intro:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3576
      then show ?thesis unfolding not_essential_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3577
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3578
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3579
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3580
  ultimately show ?thesis using f_ness unfolding not_essential_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3581
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3582
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3583
lemma isolated_singularity_at_inverse[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3584
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3585
      and f_ness:"not_essential f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3586
  shows "isolated_singularity_at (\<lambda>w. inverse (f w)) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3587
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3588
  define vf where "vf = (\<lambda>w. inverse (f w))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3589
  have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3590
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3591
    have "\<forall>\<^sub>Fw in (at z). f w=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3592
      using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3593
    then have "\<forall>\<^sub>Fw in (at z). vf w=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3594
      unfolding vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3595
    then obtain d1 where "d1>0" and d1:"\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> vf x = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3596
      unfolding eventually_at by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3597
    then have "vf holomorphic_on ball z d1-{z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3598
      apply (rule_tac holomorphic_transform[of "\<lambda>_. 0"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3599
      by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3600
    then have "vf analytic_on ball z d1 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3601
      by (simp add: analytic_on_open open_delete)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3602
    then show ?thesis using \<open>d1>0\<close> unfolding isolated_singularity_at_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3603
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3604
  moreover have ?thesis when f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3605
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3606
    have "\<forall>\<^sub>F w in at z. f w \<noteq> 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3607
    then obtain d1 where d1:"d1>0" "\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> f x \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3608
      unfolding eventually_at by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3609
    obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3610
      using f_iso unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3611
    define d3 where "d3=min d1 d2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3612
    have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3613
    moreover have "vf analytic_on ball z d3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3614
      unfolding vf_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3615
      apply (rule analytic_on_inverse)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3616
      subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3617
      subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3618
      done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3619
    ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3620
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3621
  ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3622
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3623
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3624
lemma not_essential_divide[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3625
  assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3626
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3627
  shows "not_essential (\<lambda>w. f w / g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3628
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3629
  have "not_essential (\<lambda>w. f w * inverse (g w)) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3630
    apply (rule not_essential_times[where g="\<lambda>w. inverse (g w)"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3631
    using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3632
  then show ?thesis by (simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3633
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3634
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3635
lemma 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3636
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3637
      and g_iso:"isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3638
    shows isolated_singularity_at_times[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3639
              "isolated_singularity_at (\<lambda>w. f w * g w) z" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3640
          isolated_singularity_at_add[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3641
              "isolated_singularity_at (\<lambda>w. f w + g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3642
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3643
  obtain d1 d2 where "d1>0" "d2>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3644
      and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3645
    using f_iso g_iso unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3646
  define d3 where "d3=min d1 d2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3647
  have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3648
  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3649
  have "(\<lambda>w. f w * g w) analytic_on ball z d3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3650
    apply (rule analytic_on_mult)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3651
    using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3652
  then show "isolated_singularity_at (\<lambda>w. f w * g w) z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3653
    using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3654
  have "(\<lambda>w. f w + g w) analytic_on ball z d3 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3655
    apply (rule analytic_on_add)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3656
    using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3657
  then show "isolated_singularity_at (\<lambda>w. f w + g w) z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3658
    using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3659
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3660
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3661
lemma isolated_singularity_at_uminus[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3662
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3663
  shows "isolated_singularity_at (\<lambda>w. - f w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3664
  using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3665
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3666
lemma isolated_singularity_at_id[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3667
     "isolated_singularity_at (\<lambda>w. w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3668
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3669
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3670
lemma isolated_singularity_at_minus[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3671
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3672
      and g_iso:"isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3673
    shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3674
  using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\<lambda>w. - g w"]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3675
        ,OF g_iso] by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3676
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3677
lemma isolated_singularity_at_divide[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3678
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3679
      and g_iso:"isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3680
      and g_ness:"not_essential g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3681
    shows "isolated_singularity_at (\<lambda>w. f w / g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3682
  using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso,
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3683
          of "\<lambda>w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3684
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3685
lemma isolated_singularity_at_const[singularity_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3686
    "isolated_singularity_at (\<lambda>w. c) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3687
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3688
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3689
lemma isolated_singularity_at_holomorphic:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3690
  assumes "f holomorphic_on s-{z}" "open s" "z\<in>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3691
  shows "isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3692
  using assms unfolding isolated_singularity_at_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3693
  by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3694
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3695
subsubsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3696
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3697
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
  3698
definition\<^marker>\<open>tag important\<close> zorder :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> int" where
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3699
  "zorder f z = (THE n. (\<exists>h r. r>0 \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3700
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w-z) powr (of_int n)
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3701
                   \<and> h w \<noteq>0)))"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3702
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 70113
diff changeset
  3703
definition\<^marker>\<open>tag important\<close> zor_poly
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3704
    ::"[complex \<Rightarrow> complex, complex] \<Rightarrow> complex \<Rightarrow> complex" where
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3705
  "zor_poly f z = (SOME h. \<exists>r. r > 0 \<and> h holomorphic_on cball z r \<and> h z \<noteq> 0
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3706
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w - z) powr (zorder f z)
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  3707
                   \<and> h w \<noteq>0))"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3708
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3709
lemma zorder_exist:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3710
  fixes f::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3711
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3712
  assumes f_iso:"isolated_singularity_at f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3713
      and f_ness:"not_essential f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3714
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3715
  shows "g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> g holomorphic_on cball z r
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3716
    \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w * (w-z) powr n  \<and> g w \<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3717
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3718
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3719
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3720
  have "\<exists>!n. \<exists>g r. P n g r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3721
    using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3722
  then have "\<exists>g r. P n g r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3723
    unfolding n_def P_def zorder_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3724
    by (drule_tac theI',argo)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3725
  then have "\<exists>r. P n g r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3726
    unfolding P_def zor_poly_def g_def n_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3727
    by (drule_tac someI_ex,argo)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3728
  then obtain r1 where "P n g r1" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3729
  then show ?thesis unfolding P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3730
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3731
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3732
lemma 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3733
  fixes f::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3734
  assumes f_iso:"isolated_singularity_at f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3735
      and f_ness:"not_essential f z"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3736
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3737
    shows zorder_inverse: "zorder (\<lambda>w. inverse (f w)) z = - zorder f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3738
      and zor_poly_inverse: "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. inverse (f w)) z w 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3739
                                                = inverse (zor_poly f z w)"
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  3740
proof -
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3741
  define vf where "vf = (\<lambda>w. inverse (f w))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3742
  define fn vfn where 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3743
    "fn = zorder f z"  and "vfn = zorder vf z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3744
  define fp vfp where 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3745
    "fp = zor_poly f z" and "vfp = zor_poly vf z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3746
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3747
  obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3748
          and fr: "fp holomorphic_on cball z fr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3749
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3750
    using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3751
    by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3752
  have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3753
        and fr_nz: "inverse (fp w)\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3754
    when "w\<in>ball z fr - {z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3755
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3756
    have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3757
      using fr(2)[rule_format,of w] that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3758
    then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3759
      unfolding vf_def by (auto simp add:powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3760
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3761
  obtain vfr where [simp]:"vfp z \<noteq> 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3762
      "(\<forall>w\<in>cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \<and> vfp w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3763
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3764
    have "isolated_singularity_at vf z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3765
      using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3766
    moreover have "not_essential vf z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3767
      using not_essential_inverse[OF f_ness f_iso] unfolding vf_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3768
    moreover have "\<exists>\<^sub>F w in at z. vf w \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3769
      using f_nconst unfolding vf_def by (auto elim:frequently_elim1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3770
    ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3771
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3772
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3773
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3774
  define r1 where "r1 = min fr vfr"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3775
  have "r1>0" using \<open>fr>0\<close> \<open>vfr>0\<close> unfolding r1_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3776
  show "vfn = - fn"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3777
    apply (rule holomorphic_factor_unique[of r1 vfp z "\<lambda>w. inverse (fp w)" vf])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3778
    subgoal using \<open>r1>0\<close> by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3779
    subgoal by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3780
    subgoal by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3781
    subgoal
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3782
    proof (rule ballI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3783
      fix w assume "w \<in> ball z r1 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3784
      then have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3785
      from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3786
      show "vf w = vfp w * (w - z) powr of_int vfn \<and> vfp w \<noteq> 0 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3787
              \<and> vf w = inverse (fp w) * (w - z) powr of_int (- fn) \<and> inverse (fp w) \<noteq> 0" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3788
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3789
    subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3790
    subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3791
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3792
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3793
  have "vfp w = inverse (fp w)" when "w\<in>ball z r1-{z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3794
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3795
    have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  "w\<noteq>z" using that unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3796
    from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \<open>vfn = - fn\<close> \<open>w\<noteq>z\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3797
    show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3798
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3799
  then show "\<forall>\<^sub>Fw in (at z). vfp w = inverse (fp w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3800
    unfolding eventually_at using \<open>r1>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3801
    apply (rule_tac x=r1 in exI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3802
    by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3803
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3804
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3805
lemma 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3806
  fixes f g::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3807
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3808
      and f_ness:"not_essential f z" and g_ness:"not_essential g z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3809
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3810
  shows zorder_times:"zorder (\<lambda>w. f w * g w) z = zorder f z + zorder g z" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3811
        zor_poly_times:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w * g w) z w 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3812
                                                  = zor_poly f z w *zor_poly g z w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3813
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3814
  define fg where "fg = (\<lambda>w. f w * g w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3815
  define fn gn fgn where 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3816
    "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3817
  define fp gp fgp where 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3818
    "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3819
  have f_nconst:"\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst:"\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3820
    using fg_nconst by (auto elim!:frequently_elim1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3821
  obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3822
          and fr: "fp holomorphic_on cball z fr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3823
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3824
    using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3825
  obtain gr where [simp]:"gp z \<noteq> 0" and "gr > 0"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3826
          and gr: "gp holomorphic_on cball z gr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3827
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3828
    using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3829
  define r1 where "r1=min fr gr"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3830
  have "r1>0" unfolding r1_def using \<open>fr>0\<close> \<open>gr>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3831
  have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3832
    when "w\<in>ball z r1 - {z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3833
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3834
    have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3835
      using fr(2)[rule_format,of w] that unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3836
    moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3837
      using gr(2)[rule_format, of w] that unfolding r1_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3838
    ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3839
      unfolding fg_def by (auto simp add:powr_add)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3840
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3841
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3842
  obtain fgr where [simp]:"fgp z \<noteq> 0" and "fgr > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3843
          and fgr: "fgp holomorphic_on cball z fgr" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3844
                  "\<forall>w\<in>cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3845
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3846
    have "fgp z \<noteq> 0 \<and> (\<exists>r>0. fgp holomorphic_on cball z r 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3847
            \<and> (\<forall>w\<in>cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3848
      apply (rule zorder_exist[of fg z, folded fgn_def fgp_def])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3849
      subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3850
      subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3851
      subgoal unfolding fg_def using fg_nconst .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3852
      done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3853
    then show ?thesis using that by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3854
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3855
  define r2 where "r2 = min fgr r1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3856
  have "r2>0" using \<open>r1>0\<close> \<open>fgr>0\<close> unfolding r2_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3857
  show "fgn = fn + gn "
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3858
    apply (rule holomorphic_factor_unique[of r2 fgp z "\<lambda>w. fp w * gp w" fg])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3859
    subgoal using \<open>r2>0\<close> by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3860
    subgoal by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3861
    subgoal by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3862
    subgoal
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3863
    proof (rule ballI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3864
      fix w assume "w \<in> ball z r2 - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3865
      then have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}"  unfolding r2_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3866
      from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3867
      show "fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3868
              \<and> fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \<and> fp w * gp w \<noteq> 0" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3869
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3870
    subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3871
    subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3872
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3873
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3874
  have "fgp w = fp w *gp w" when "w\<in>ball z r2-{z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3875
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3876
    have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}" "w\<noteq>z" using that  unfolding r2_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3877
    from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \<open>fgn = fn + gn\<close> \<open>w\<noteq>z\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3878
    show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3879
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3880
  then show "\<forall>\<^sub>Fw in (at z). fgp w = fp w * gp w" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3881
    using \<open>r2>0\<close> unfolding eventually_at by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3882
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3883
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3884
lemma 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3885
  fixes f g::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3886
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3887
      and f_ness:"not_essential f z" and g_ness:"not_essential g z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3888
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3889
  shows zorder_divide:"zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3890
        zor_poly_divide:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3891
                                                  = zor_poly f z w  / zor_poly g z w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3892
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3893
  have f_nconst:"\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst:"\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3894
    using fg_nconst by (auto elim!:frequently_elim1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3895
  define vg where "vg=(\<lambda>w. inverse (g w))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3896
  have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3897
    apply (rule zorder_times[OF f_iso _ f_ness,of vg])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3898
    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3899
    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3900
    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3901
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3902
  then show "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3903
    using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3904
    by (auto simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3905
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3906
  have "\<forall>\<^sub>F w in at z. zor_poly (\<lambda>w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3907
    apply (rule zor_poly_times[OF f_iso _ f_ness,of vg])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3908
    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3909
    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3910
    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3911
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3912
  then show "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w = zor_poly f z w  / zor_poly g z w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3913
    using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3914
    apply eventually_elim
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3915
    by (auto simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3916
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3917
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3918
lemma zorder_exist_zero:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3919
  fixes f::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3920
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3921
  assumes  holo: "f holomorphic_on s" and 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3922
          "open s" "connected s" "z\<in>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3923
      and non_const: "\<exists>w\<in>s. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3924
  shows "(if f z=0 then n > 0 else n=0) \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> s \<and> g holomorphic_on cball z r
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3925
    \<and> (\<forall>w\<in>cball z r. f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3926
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3927
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3928
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3929
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3930
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3931
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3932
    proof (rule zorder_exist[of f z,folded g_def n_def])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3933
      show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3934
        using holo assms(4,6)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3935
        by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3936
      show "not_essential f z" unfolding not_essential_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3937
        using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3938
        by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3939
      have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w\<in>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3940
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3941
        obtain w where "w\<in>s" "f w\<noteq>0" using non_const by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3942
        then show ?thesis 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3943
          by (rule non_zero_neighbour_alt[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3944
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3945
      then show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3946
        apply (elim eventually_frequentlyE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3947
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3948
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3949
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3950
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3951
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3952
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3953
      using assms(4,6) open_contains_cball_eq by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3954
    define r3 where "r3=min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3955
    have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3956
    moreover have "g holomorphic_on cball z r3" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3957
      using r1(1) unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3958
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3959
      using r1(2) unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3960
    ultimately show ?thesis using that[of r3] \<open>g z\<noteq>0\<close> by auto 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3961
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3962
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3963
  have if_0:"if f z=0 then n > 0 else n=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3964
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3965
    have "f\<midarrow> z \<rightarrow> f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3966
      by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3967
    then have "(\<lambda>w. g w * (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3968
      apply (elim Lim_transform_within_open[where s="ball z r"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3969
      using r by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3970
    moreover have "g \<midarrow>z\<rightarrow>g z"
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  3971
      by (metis (mono_tags, lifting) open_ball at_within_open_subset 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3972
          ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3973
    ultimately have "(\<lambda>w. (g w * (w - z) powr of_int n) / g w) \<midarrow>z\<rightarrow> f z/g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3974
      apply (rule_tac tendsto_divide)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3975
      using \<open>g z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3976
    then have powr_tendsto:"(\<lambda>w. (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z/g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3977
      apply (elim Lim_transform_within_open[where s="ball z r"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3978
      using r by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3979
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3980
    have ?thesis when "n\<ge>0" "f z=0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3981
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3982
      have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3983
        using powr_tendsto 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3984
        apply (elim Lim_transform_within[where d=r])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3985
        by (auto simp add: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3986
      then have *:"(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>f z=0\<close> by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3987
      moreover have False when "n=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3988
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3989
        have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3990
          using \<open>n=0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3991
        then show False using * using LIM_unique zero_neq_one by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3992
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3993
      ultimately show ?thesis using that by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3994
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3995
    moreover have ?thesis when "n\<ge>0" "f z\<noteq>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3996
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3997
      have False when "n>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3998
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  3999
        have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4000
          using powr_tendsto 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4001
          apply (elim Lim_transform_within[where d=r])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4002
          by (auto simp add: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4003
        moreover have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4004
          using \<open>n>0\<close> by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4005
        ultimately show False using \<open>f z\<noteq>0\<close> \<open>g z\<noteq>0\<close> using LIM_unique divide_eq_0_iff by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4006
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4007
      then show ?thesis using that by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4008
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4009
    moreover have False when "n<0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4010
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4011
      have "(\<lambda>w. inverse ((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> f z/g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4012
           "(\<lambda>w.((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4013
        subgoal  using powr_tendsto powr_of_int that
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4014
          by (elim Lim_transform_within_open[where s=UNIV],auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4015
        subgoal using that by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4016
        done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4017
      from tendsto_mult[OF this,simplified] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4018
      have "(\<lambda>x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \<midarrow>z\<rightarrow> 0" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4019
      then have "(\<lambda>x. 1::complex) \<midarrow>z\<rightarrow> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4020
        by (elim Lim_transform_within_open[where s=UNIV],auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4021
      then show False using LIM_const_eq by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4022
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4023
    ultimately show ?thesis by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4024
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4025
  moreover have "f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0" when "w\<in>cball z r" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4026
  proof (cases "w=z")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4027
    case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4028
    then have "f \<midarrow>z\<rightarrow>f w" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4029
      using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4030
    then have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow>f w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4031
    proof (elim Lim_transform_within[OF _ \<open>r>0\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4032
      fix x assume "0 < dist x z" "dist x z < r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4033
      then have "x \<in> cball z r - {z}" "x\<noteq>z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4034
        unfolding cball_def by (auto simp add: dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4035
      then have "f x = g x * (x - z) powr of_int n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4036
        using r(4)[rule_format,of x] by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4037
      also have "... = g x * (x - z) ^ nat n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4038
        apply (subst powr_of_int)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4039
        using if_0 \<open>x\<noteq>z\<close> by (auto split:if_splits)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4040
      finally show "f x = g x * (x - z) ^ nat n" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4041
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4042
    moreover have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow> g w * (w-z) ^ nat n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4043
      using True apply (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4044
      by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4045
          continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4046
    ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4047
    then show ?thesis using \<open>g z\<noteq>0\<close> True by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4048
  next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4049
    case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4050
    then have "f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4051
      using r(4) that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4052
    then show ?thesis using False if_0 powr_of_int by (auto split:if_splits)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4053
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4054
  ultimately show ?thesis using r by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4055
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4056
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4057
lemma zorder_exist_pole:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4058
  fixes f::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4059
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4060
  assumes  holo: "f holomorphic_on s-{z}" and 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4061
          "open s" "z\<in>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4062
      and "is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4063
  shows "n < 0 \<and> g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> s \<and> g holomorphic_on cball z r
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4064
    \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4065
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4066
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4067
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4068
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4069
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4070
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4071
    proof (rule zorder_exist[of f z,folded g_def n_def])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4072
      show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4073
        using holo assms(4,5)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4074
        by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4075
      show "not_essential f z" unfolding not_essential_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4076
        using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4077
        by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4078
      from non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4079
        apply (elim eventually_frequentlyE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4080
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4081
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4082
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4083
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4084
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4085
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4086
      using assms(4,5) open_contains_cball_eq by metis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4087
    define r3 where "r3=min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4088
    have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4089
    moreover have "g holomorphic_on cball z r3" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4090
      using r1(1) unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4091
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4092
      using r1(2) unfolding r3_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4093
    ultimately show ?thesis using that[of r3] \<open>g z\<noteq>0\<close> by auto 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4094
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4095
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4096
  have "n<0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4097
  proof (rule ccontr)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4098
    assume " \<not> n < 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4099
    define c where "c=(if n=0 then g z else 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4100
    have [simp]:"g \<midarrow>z\<rightarrow> g z" 
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  4101
      by (metis open_ball at_within_open ball_subset_cball centre_in_ball 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4102
            continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) )
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4103
    have "\<forall>\<^sub>F x in at z. f x = g x * (x - z) ^ nat n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4104
      unfolding eventually_at_topological
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4105
      apply (rule_tac exI[where x="ball z r"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4106
      using r powr_of_int \<open>\<not> n < 0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4107
    moreover have "(\<lambda>x. g x * (x - z) ^ nat n) \<midarrow>z\<rightarrow>c"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4108
    proof (cases "n=0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4109
      case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4110
      then show ?thesis unfolding c_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4111
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4112
      case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4113
      then have "(\<lambda>x. (x - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>\<not> n < 0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4114
        by (auto intro!:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4115
      from tendsto_mult[OF _ this,of g "g z",simplified] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4116
      show ?thesis unfolding c_def using False by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4117
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4118
    ultimately have "f \<midarrow>z\<rightarrow>c" using tendsto_cong by fast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4119
    then show False using \<open>is_pole f z\<close> at_neq_bot not_tendsto_and_filterlim_at_infinity 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4120
      unfolding is_pole_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4121
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4122
  moreover have "\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4123
    using r(4) \<open>n<0\<close> powr_of_int 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4124
    by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4125
  ultimately show ?thesis using r(1-3) \<open>g z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4126
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4127
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4128
lemma zorder_eqI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4129
  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4130
  assumes fg_eq:"\<And>w. \<lbrakk>w \<in> s;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w - z) powr n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4131
  shows   "zorder f z = n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4132
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4133
  have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4134
  moreover have "open (-{0::complex})" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4135
  ultimately have "open ((g -` (-{0})) \<inter> s)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4136
    unfolding continuous_on_open_vimage[OF \<open>open s\<close>] by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4137
  moreover from assms have "z \<in> (g -` (-{0})) \<inter> s" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4138
  ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  s \<inter> (g -` (-{0}))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4139
    unfolding open_contains_cball by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4140
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4141
  let ?gg= "(\<lambda>w. g w * (w - z) powr n)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4142
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4143
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4144
  have "P n g r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4145
    unfolding P_def using r assms(3,4,5) by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4146
  then have "\<exists>g r. P n g r" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4147
  moreover have unique: "\<exists>!n. \<exists>g r. P n g r" unfolding P_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4148
  proof (rule holomorphic_factor_puncture)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4149
    have "ball z r-{z} \<subseteq> s" using r using ball_subset_cball by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4150
    then have "?gg holomorphic_on ball z r-{z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4151
      using \<open>g holomorphic_on s\<close> r by (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4152
    then have "f holomorphic_on ball z r - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4153
      apply (elim holomorphic_transform)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4154
      using fg_eq \<open>ball z r-{z} \<subseteq> s\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4155
    then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4156
      using analytic_on_open open_delete r(1) by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4157
  next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4158
    have "not_essential ?gg z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4159
    proof (intro singularity_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4160
      show "not_essential g z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4161
        by (meson \<open>continuous_on s g\<close> assms(1) assms(2) continuous_on_eq_continuous_at 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4162
            isCont_def not_essential_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4163
      show " \<forall>\<^sub>F w in at z. w - z \<noteq> 0" by (simp add: eventually_at_filter)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4164
      then show "LIM w at z. w - z :> at 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4165
        unfolding filterlim_at by (auto intro:tendsto_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4166
      show "isolated_singularity_at g z" 
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  4167
        by (meson Diff_subset open_ball analytic_on_holomorphic 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4168
            assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4169
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4170
    then show "not_essential f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4171
      apply (elim not_essential_transform)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4172
      unfolding eventually_at using assms(1,2) assms(5)[symmetric] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4173
      by (metis dist_commute mem_ball openE subsetCE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4174
    show "\<exists>\<^sub>F w in at z. f w \<noteq> 0" unfolding frequently_at 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4175
    proof (rule,rule)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4176
      fix d::real assume "0 < d"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4177
      define z' where "z'=z+min d r / 2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4178
      have "z' \<noteq> z" " dist z' z < d "
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4179
        unfolding z'_def using \<open>d>0\<close> \<open>r>0\<close> 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4180
        by (auto simp add:dist_norm)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4181
      moreover have "f z' \<noteq> 0"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4182
      proof (subst fg_eq[OF _ \<open>z'\<noteq>z\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4183
        have "z' \<in> cball z r" unfolding z'_def using \<open>r>0\<close> \<open>d>0\<close> by (auto simp add:dist_norm)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4184
        then show " z' \<in> s" using r(2) by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4185
        show "g z' * (z' - z) powr of_int n \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4186
          using P_def \<open>P n g r\<close> \<open>z' \<in> cball z r\<close> calculation(1) by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4187
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4188
      ultimately show "\<exists>x\<in>UNIV. x \<noteq> z \<and> dist x z < d \<and> f x \<noteq> 0" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4189
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4190
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4191
  ultimately have "(THE n. \<exists>g r. P n g r) = n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4192
    by (rule_tac the1_equality)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4193
  then show ?thesis unfolding zorder_def P_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4194
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4195
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4196
lemma residue_pole_order:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4197
  fixes f::"complex \<Rightarrow> complex" and z::complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4198
  defines "n \<equiv> nat (- zorder f z)" and "h \<equiv> zor_poly f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4199
  assumes f_iso:"isolated_singularity_at f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4200
    and pole:"is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4201
  shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4202
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4203
  define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4204
  obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4205
    using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4206
  obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> ball z e" and h_holo: "h holomorphic_on cball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4207
      and h_divide:"(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4208
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4209
    obtain r where r:"zorder f z < 0" "h z \<noteq> 0" "r>0" "cball z r \<subseteq> ball z e" "h holomorphic_on cball z r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4210
        "(\<forall>w\<in>cball z r - {z}. f w = h w / (w - z) ^ n \<and> h w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4211
      using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>,folded n_def h_def] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4212
    have "n>0" using \<open>zorder f z < 0\<close> unfolding n_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4213
    moreover have "(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4214
      using \<open>h z\<noteq>0\<close> r(6) by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4215
    ultimately show ?thesis using r(3,4,5) that by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4216
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4217
  have r_nonzero:"\<And>w. w \<in> ball z r - {z} \<Longrightarrow> f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4218
    using h_divide by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4219
  define c where "c \<equiv> 2 * pi * \<i>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4220
  define der_f where "der_f \<equiv> ((deriv ^^ (n - 1)) h z / fact (n-1))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4221
  define h' where "h' \<equiv> \<lambda>u. h u / (u - z) ^ n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4222
  have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4223
    unfolding h'_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4224
    proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1",
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4225
        folded c_def Suc_pred'[OF \<open>n>0\<close>]])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4226
      show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4227
      show "h holomorphic_on ball z r" using h_holo by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4228
      show " z \<in> ball z r" using \<open>r>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4229
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4230
  then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4231
  then have "(f has_contour_integral c * der_f) (circlepath z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4232
    proof (elim has_contour_integral_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4233
      fix x assume "x \<in> path_image (circlepath z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4234
      hence "x\<in>cball z r - {z}" using \<open>r>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4235
      then show "h' x = f x" using h_divide unfolding h'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4236
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4237
  moreover have "(f has_contour_integral c * residue f z) (circlepath z r)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4238
    using base_residue[of \<open>ball z e\<close> z,simplified,OF \<open>r>0\<close> f_holo r_cball,folded c_def] 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4239
    unfolding c_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4240
  ultimately have "c * der_f =  c * residue f z" using has_contour_integral_unique by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4241
  hence "der_f = residue f z" unfolding c_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4242
  thus ?thesis unfolding der_f_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4243
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4244
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4245
lemma simple_zeroI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4246
  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4247
  assumes "\<And>w. w \<in> s \<Longrightarrow> f w = g w * (w - z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4248
  shows   "zorder f z = 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4249
  using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4250
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4251
lemma higher_deriv_power:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4252
  shows   "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) w = 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4253
             pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4254
proof (induction j arbitrary: w)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4255
  case 0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4256
  thus ?case by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4257
next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4258
  case (Suc j w)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4259
  have "(deriv ^^ Suc j) (\<lambda>w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\<lambda>w. (w - z) ^ n)) w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4260
    by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4261
  also have "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) = 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4262
               (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4263
    using Suc by (intro Suc.IH ext)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4264
  also {
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4265
    have "(\<dots> has_field_derivative of_nat (n - j) *
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4266
               pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4267
      using Suc.prems by (auto intro!: derivative_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4268
    also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4269
                 pochhammer (of_nat (Suc n - Suc j)) (Suc j)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4270
      by (cases "Suc j \<le> n", subst pochhammer_rec) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4271
         (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4272
    finally have "deriv (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w =
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4273
                    \<dots> * (w - z) ^ (n - Suc j)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4274
      by (rule DERIV_imp_deriv)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4275
  }
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4276
  finally show ?case .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4277
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4278
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4279
lemma zorder_zero_eqI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4280
  assumes  f_holo:"f holomorphic_on s" and "open s" "z \<in> s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4281
  assumes zero: "\<And>i. i < nat n \<Longrightarrow> (deriv ^^ i) f z = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4282
  assumes nz: "(deriv ^^ nat n) f z \<noteq> 0" and "n\<ge>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4283
  shows   "zorder f z = n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4284
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4285
  obtain r where [simp]:"r>0" and "ball z r \<subseteq> s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4286
    using \<open>open s\<close> \<open>z\<in>s\<close> openE by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4287
  have nz':"\<exists>w\<in>ball z r. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4288
  proof (rule ccontr)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4289
    assume "\<not> (\<exists>w\<in>ball z r. f w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4290
    then have "eventually (\<lambda>u. f u = 0) (nhds z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4291
      using \<open>r>0\<close> unfolding eventually_nhds 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4292
      apply (rule_tac x="ball z r" in exI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4293
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4294
    then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\<lambda>_. 0) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4295
      by (intro higher_deriv_cong_ev) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4296
    also have "(deriv ^^ nat n) (\<lambda>_. 0) z = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4297
      by (induction n) simp_all
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4298
    finally show False using nz by contradiction
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4299
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4300
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4301
  define zn g where "zn = zorder f z" and "g = zor_poly f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4302
  obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4303
            [simp]:"e>0" and "cball z e \<subseteq> ball z r" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4304
            g_holo:"g holomorphic_on cball z e" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4305
            e_fac:"(\<forall>w\<in>cball z e. f w = g w * (w - z) ^ nat zn \<and> g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4306
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4307
    have "f holomorphic_on ball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4308
      using f_holo \<open>ball z r \<subseteq> s\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4309
    from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4310
    show ?thesis by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4311
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4312
  from this(1,2,5) have "zn\<ge>0" "g z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4313
    subgoal by (auto split:if_splits) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4314
    subgoal using \<open>0 < e\<close> ball_subset_cball centre_in_ball e_fac by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4315
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4316
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4317
  define A where "A = (\<lambda>i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4318
  have deriv_A:"(deriv ^^ i) f z = (if zn \<le> int i then A i else 0)" for i
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4319
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4320
    have "eventually (\<lambda>w. w \<in> ball z e) (nhds z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4321
      using \<open>cball z e \<subseteq> ball z r\<close> \<open>e>0\<close> by (intro eventually_nhds_in_open) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4322
    hence "eventually (\<lambda>w. f w = (w - z) ^ (nat zn) * g w) (nhds z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4323
      apply eventually_elim 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4324
      by (use e_fac in auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4325
    hence "(deriv ^^ i) f z = (deriv ^^ i) (\<lambda>w. (w - z) ^ nat zn * g w) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4326
      by (intro higher_deriv_cong_ev) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4327
    also have "\<dots> = (\<Sum>j=0..i. of_nat (i choose j) *
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4328
                       (deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4329
      using g_holo \<open>e>0\<close> 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4330
      by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4331
    also have "\<dots> = (\<Sum>j=0..i. if j = nat zn then 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4332
                    of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4333
    proof (intro sum.cong refl, goal_cases)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4334
      case (1 j)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4335
      have "(deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z = 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4336
              pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4337
        by (subst higher_deriv_power) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4338
      also have "\<dots> = (if j = nat zn then fact j else 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4339
        by (auto simp: not_less pochhammer_0_left pochhammer_fact)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4340
      also have "of_nat (i choose j) * \<dots> * (deriv ^^ (i - j)) g z = 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4341
                   (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4342
                        * (deriv ^^ (i - nat zn)) g z else 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4343
        by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4344
      finally show ?case .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4345
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4346
    also have "\<dots> = (if i \<ge> zn then A i else 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4347
      by (auto simp: A_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4348
    finally show "(deriv ^^ i) f z = \<dots>" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4349
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4350
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4351
  have False when "n<zn"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4352
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4353
    have "(deriv ^^ nat n) f z = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4354
      using deriv_A[of "nat n"] that \<open>n\<ge>0\<close> by auto 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4355
    with nz show False by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4356
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4357
  moreover have "n\<le>zn"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4358
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4359
    have "g z \<noteq> 0" using e_fac[rule_format,of z] \<open>e>0\<close> by simp 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4360
    then have "(deriv ^^ nat zn) f z \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4361
      using deriv_A[of "nat zn"] by(auto simp add:A_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4362
    then have "nat zn \<ge> nat n" using zero[of "nat zn"] by linarith
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4363
    moreover have "zn\<ge>0" using e_if by (auto split:if_splits)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4364
    ultimately show ?thesis using nat_le_eq_zle by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4365
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4366
  ultimately show ?thesis unfolding zn_def by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4367
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4368
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4369
lemma 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4370
  assumes "eventually (\<lambda>z. f z = g z) (at z)" "z = z'"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4371
  shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4372
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4373
  define P where "P = (\<lambda>ff n h r. 0 < r \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4374
                    \<and> (\<forall>w\<in>cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \<and> h w\<noteq>0))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4375
  have "(\<exists>r. P f n h r) = (\<exists>r. P g n h r)" for n h 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4376
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4377
    have *: "\<exists>r. P g n h r" if "\<exists>r. P f n h r" and "eventually (\<lambda>x. f x = g x) (at z)" for f g 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4378
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4379
      from that(1) obtain r1 where r1_P:"P f n h r1" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4380
      from that(2) obtain r2 where "r2>0" and r2_dist:"\<forall>x. x \<noteq> z \<and> dist x z \<le> r2 \<longrightarrow> f x = g x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4381
        unfolding eventually_at_le by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4382
      define r where "r=min r1 r2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4383
      have "r>0" "h z\<noteq>0" using r1_P \<open>r2>0\<close> unfolding r_def P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4384
      moreover have "h holomorphic_on cball z r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4385
        using r1_P unfolding P_def r_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4386
      moreover have "g w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0" when "w\<in>cball z r - {z}" for w
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4387
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4388
        have "f w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4389
          using r1_P that unfolding P_def r_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4390
        moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4391
          by (simp add: dist_commute) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4392
        ultimately show ?thesis by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4393
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4394
      ultimately show ?thesis unfolding P_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4395
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4396
    from assms have eq': "eventually (\<lambda>z. g z = f z) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4397
      by (simp add: eq_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4398
    show ?thesis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4399
      by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4400
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4401
  then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4402
      using \<open>z=z'\<close> unfolding P_def zorder_def zor_poly_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4403
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4404
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4405
lemma zorder_nonzero_div_power:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4406
  assumes "open s" "z \<in> s" "f holomorphic_on s" "f z \<noteq> 0" "n > 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4407
  shows  "zorder (\<lambda>w. f w / (w - z) ^ n) z = - n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4408
  apply (rule zorder_eqI[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>f holomorphic_on s\<close> \<open>f z\<noteq>0\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4409
  apply (subst powr_of_int)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4410
  using \<open>n>0\<close> by (auto simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4411
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4412
lemma zor_poly_eq:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4413
  assumes "isolated_singularity_at f z" "not_essential f z" "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4414
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4415
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4416
  obtain r where r:"r>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4417
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4418
    using zorder_exist[OF assms] by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4419
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4420
    by (auto simp: field_simps powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4421
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4422
    using r eventually_at_ball'[of r z UNIV] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4423
  thus ?thesis by eventually_elim (insert *, auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4424
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4425
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4426
lemma zor_poly_zero_eq:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4427
  assumes "f holomorphic_on s" "open s" "connected s" "z \<in> s" "\<exists>w\<in>s. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4428
  shows "eventually (\<lambda>w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4429
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4430
  obtain r where r:"r>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4431
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4432
    using zorder_exist_zero[OF assms] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4433
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4434
    by (auto simp: field_simps powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4435
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4436
    using r eventually_at_ball'[of r z UNIV] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4437
  thus ?thesis by eventually_elim (insert *, auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4438
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4439
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4440
lemma zor_poly_pole_eq:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4441
  assumes f_iso:"isolated_singularity_at f z" "is_pole f z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4442
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4443
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4444
  obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4445
    using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4446
  obtain r where r:"r>0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4447
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4448
    using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4449
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" 
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4450
    by (auto simp: field_simps)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4451
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4452
    using r eventually_at_ball'[of r z UNIV] by auto
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4453
  thus ?thesis by eventually_elim (insert *, auto)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4454
qed
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4455
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4456
lemma zor_poly_eqI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4457
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4458
  defines "n \<equiv> zorder f z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4459
  assumes "isolated_singularity_at f z0" "not_essential f z0" "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4460
  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> c) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4461
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4462
  shows   "zor_poly f z0 z0 = c"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4463
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4464
  from zorder_exist[OF assms(2-4)] obtain r where
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4465
    r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4466
       "\<And>w. w \<in> cball z0 r - {z0} \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) powr n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4467
    unfolding n_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4468
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4469
    using eventually_at_ball'[of r z0 UNIV] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4470
  hence "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4471
    by eventually_elim (insert r, auto simp: field_simps powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4472
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4473
    using r by (intro holomorphic_on_imp_continuous_on) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4474
  with r(1,2) have "isCont (zor_poly f z0) z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4475
    by (auto simp: continuous_on_eq_continuous_at)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4476
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4477
    unfolding isCont_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4478
  ultimately have "((\<lambda>w. f w * (w - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4479
    by (rule Lim_transform_eventually)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4480
  hence "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4481
    by (rule filterlim_compose[OF _ g])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4482
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4483
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4484
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4485
lemma zor_poly_zero_eqI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4486
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4487
  defines "n \<equiv> zorder f z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4488
  assumes "f holomorphic_on A" "open A" "connected A" "z0 \<in> A" "\<exists>z\<in>A. f z \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4489
  assumes lim: "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> c) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4490
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4491
  shows   "zor_poly f z0 z0 = c"
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4492
proof -
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4493
  from zorder_exist_zero[OF assms(2-6)] obtain r where
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4494
    r: "r > 0" "cball z0 r \<subseteq> A" "zor_poly f z0 holomorphic_on cball z0 r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4495
       "\<And>w. w \<in> cball z0 r \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) ^ nat n"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4496
    unfolding n_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4497
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4498
    using eventually_at_ball'[of r z0 UNIV] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4499
  hence "eventually (\<lambda>w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4500
    by eventually_elim (insert r, auto simp: field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4501
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4502
    using r by (intro holomorphic_on_imp_continuous_on) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4503
  with r(1,2) have "isCont (zor_poly f z0) z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4504
    by (auto simp: continuous_on_eq_continuous_at)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4505
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4506
    unfolding isCont_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4507
  ultimately have "((\<lambda>w. f w / (w - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4508
    by (rule Lim_transform_eventually)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4509
  hence "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4510
    by (rule filterlim_compose[OF _ g])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4511
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4512
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4513
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4514
lemma zor_poly_pole_eqI:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4515
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4516
  defines "n \<equiv> zorder f z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4517
  assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4518
  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> c) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4519
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4520
  shows   "zor_poly f z0 z0 = c"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4521
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4522
  obtain r where r: "r > 0"  "zor_poly f z0 holomorphic_on cball z0 r"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4523
  proof -   
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4524
    have "\<exists>\<^sub>F w in at z0. f w \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4525
      using non_zero_neighbour_pole[OF \<open>is_pole f z0\<close>] by (auto elim:eventually_frequentlyE)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4526
    moreover have "not_essential f z0" unfolding not_essential_def using \<open>is_pole f z0\<close> by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4527
    ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4528
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4529
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4530
    using eventually_at_ball'[of r z0 UNIV] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4531
  have "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4532
    using zor_poly_pole_eq[OF f_iso \<open>is_pole f z0\<close>] unfolding n_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4533
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4534
    using r by (intro holomorphic_on_imp_continuous_on) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4535
  with r(1,2) have "isCont (zor_poly f z0) z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4536
    by (auto simp: continuous_on_eq_continuous_at)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4537
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4538
    unfolding isCont_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4539
  ultimately have "((\<lambda>w. f w * (w - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4540
    by (rule Lim_transform_eventually)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4541
  hence "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4542
    by (rule filterlim_compose[OF _ g])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4543
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4544
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4545
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4546
lemma residue_simple_pole:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4547
  assumes "isolated_singularity_at f z0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4548
  assumes "is_pole f z0" "zorder f z0 = - 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4549
  shows   "residue f z0 = zor_poly f z0 z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4550
  using assms by (subst residue_pole_order) simp_all
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4551
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4552
lemma residue_simple_pole_limit:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4553
  assumes "isolated_singularity_at f z0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4554
  assumes "is_pole f z0" "zorder f z0 = - 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4555
  assumes "((\<lambda>x. f (g x) * (g x - z0)) \<longlongrightarrow> c) F"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4556
  assumes "filterlim g (at z0) F" "F \<noteq> bot"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4557
  shows   "residue f z0 = c"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4558
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4559
  have "residue f z0 = zor_poly f z0 z0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4560
    by (rule residue_simple_pole assms)+
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4561
  also have "\<dots> = c"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4562
    apply (rule zor_poly_pole_eqI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4563
    using assms by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4564
  finally show ?thesis .
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4565
qed
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66456
diff changeset
  4566
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4567
lemma lhopital_complex_simple:
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4568
  assumes "(f has_field_derivative f') (at z)" 
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4569
  assumes "(g has_field_derivative g') (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4570
  assumes "f z = 0" "g z = 0" "g' \<noteq> 0" "f' / g' = c"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4571
  shows   "((\<lambda>w. f w / g w) \<longlongrightarrow> c) (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4572
proof -
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4573
  have "eventually (\<lambda>w. w \<noteq> z) (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4574
    by (auto simp: eventually_at_filter)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4575
  hence "eventually (\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4576
    by eventually_elim (simp add: assms divide_simps)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4577
  moreover have "((\<lambda>w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \<longlongrightarrow> f' / g') (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4578
    by (intro tendsto_divide has_field_derivativeD assms)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4579
  ultimately have "((\<lambda>w. f w / g w) \<longlongrightarrow> f' / g') (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4580
    by (rule Lim_transform_eventually)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4581
  with assms show ?thesis by simp
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4582
qed
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4583
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4584
lemma
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4585
  assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4586
          and "open s" "connected s" "z \<in> s" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4587
  assumes g_deriv:"(g has_field_derivative g') (at z)"
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4588
  assumes "f z \<noteq> 0" "g z = 0" "g' \<noteq> 0"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4589
  shows   porder_simple_pole_deriv: "zorder (\<lambda>w. f w / g w) z = - 1"
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4590
    and   residue_simple_pole_deriv: "residue (\<lambda>w. f w / g w) z = f z / g'"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4591
proof -
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4592
  have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4593
    using isolated_singularity_at_holomorphic[OF _ \<open>open s\<close> \<open>z\<in>s\<close>] f_holo g_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4594
    by (meson Diff_subset holomorphic_on_subset)+
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4595
  have [simp]:"not_essential f z" "not_essential g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4596
    unfolding not_essential_def using f_holo g_holo assms(3,5)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4597
    by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4598
  have g_nconst:"\<exists>\<^sub>F w in at z. g w \<noteq>0 " 
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4599
  proof (rule ccontr)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4600
    assume "\<not> (\<exists>\<^sub>F w in at z. g w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4601
    then have "\<forall>\<^sub>F w in nhds z. g w = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4602
      unfolding eventually_at eventually_nhds frequently_at using \<open>g z = 0\<close> 
69516
09bb8f470959 most of Topology_Euclidean_Space (now Elementary_Topology) requires fewer dependencies
immler
parents: 69508
diff changeset
  4603
      by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4604
    then have "deriv g z = deriv (\<lambda>_. 0) z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4605
      by (intro deriv_cong_ev) auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4606
    then have "deriv g z = 0" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4607
    then have "g' = 0" using g_deriv DERIV_imp_deriv by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4608
    then show False using \<open>g'\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4609
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4610
  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4611
  have "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4612
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4613
    have "\<forall>\<^sub>F w in at z. f w \<noteq>0 \<and> w\<in>s" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4614
      apply (rule non_zero_neighbour_alt)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4615
      using assms by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4616
    with g_nconst have "\<exists>\<^sub>F w in at z. f w * g w \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4617
      by (elim frequently_rev_mp eventually_rev_mp,auto)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4618
    then show ?thesis using zorder_divide[of f z g] by auto
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4619
  qed
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4620
  moreover have "zorder f z=0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4621
    apply (rule zorder_zero_eqI[OF f_holo \<open>open s\<close> \<open>z\<in>s\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4622
    using \<open>f z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4623
  moreover have "zorder g z=1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4624
    apply (rule zorder_zero_eqI[OF g_holo \<open>open s\<close> \<open>z\<in>s\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4625
    subgoal using assms(8) by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4626
    subgoal using DERIV_imp_deriv assms(9) g_deriv by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4627
    subgoal by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4628
    done
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4629
  ultimately show "zorder (\<lambda>w. f w / g w) z = - 1" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4630
  
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4631
  show "residue (\<lambda>w. f w / g w) z = f z / g'"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4632
  proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4633
    show "zorder (\<lambda>w. f w / g w) z = - 1" by fact
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4634
    show "isolated_singularity_at (\<lambda>w. f w / g w) z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4635
      by (auto intro: singularity_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4636
    show "is_pole (\<lambda>w. f w / g w) z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4637
    proof (rule is_pole_divide)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4638
      have "\<forall>\<^sub>F x in at z. g x \<noteq> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4639
        apply (rule non_zero_neighbour)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4640
        using g_nconst by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4641
      moreover have "g \<midarrow>z\<rightarrow> 0" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4642
        using DERIV_isCont assms(8) continuous_at g_deriv by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4643
      ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4644
      show "isCont f z" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4645
        using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4646
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4647
      show "f z \<noteq> 0" by fact
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4648
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4649
    show "filterlim id (at z) (at z)" by (simp add: filterlim_iff)
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4650
    have "((\<lambda>w. (f w * (w - z)) / g w) \<longlongrightarrow> f z / g') (at z)"
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4651
    proof (rule lhopital_complex_simple)
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4652
      show "((\<lambda>w. f w * (w - z)) has_field_derivative f z) (at z)"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4653
        using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo])
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4654
      show "(g has_field_derivative g') (at z)" by fact
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4655
    qed (insert assms, auto)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4656
    then show "((\<lambda>w. (f w / g w) * (w - z)) \<longlongrightarrow> f z / g') (at z)"
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  4657
      by (simp add: divide_simps)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4658
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4659
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4660
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4661
subsection \<open>The argument principle\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4662
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4663
theorem argument_principle:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4664
  fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set"
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
  4665
  defines "pz \<equiv> {w. f w = 0 \<or> w \<in> poles}" \<comment> \<open>\<^term>\<open>pz\<close> is the set of poles and zeros\<close>
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4666
  assumes "open s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4667
          "connected s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4668
          f_holo:"f holomorphic_on s-poles" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4669
          h_holo:"h holomorphic_on s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4670
          "valid_path g" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4671
          loop:"pathfinish g = pathstart g" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4672
          path_img:"path_image g \<subseteq> s - pz" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4673
          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4674
          finite:"finite pz" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4675
          poles:"\<forall>p\<in>poles. is_pole f p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4676
  shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> *
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4677
          (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4678
    (is "?L=?R")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4679
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4680
  define c where "c \<equiv> 2 * complex_of_real pi * \<i> "
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4681
  define ff where "ff \<equiv> (\<lambda>x. deriv f x * h x / f x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4682
  define cont where "cont \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4683
  define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4684
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4685
  have "\<exists>e>0. avoid p e \<and> (p\<in>pz \<longrightarrow> cont ff p e)" when "p\<in>s" for p
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4686
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4687
    obtain e1 where "e1>0" and e1_avoid:"avoid p e1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4688
      using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4689
    have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont ff p e2" when "p\<in>pz"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4690
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4691
      define po where "po \<equiv> zorder f p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4692
      define pp where "pp \<equiv> zor_poly f p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4693
      define f' where "f' \<equiv> \<lambda>w. pp w * (w - p) powr po"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4694
      define ff' where "ff' \<equiv> (\<lambda>x. deriv f' x * h x / f' x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4695
      obtain r where "pp p\<noteq>0" "r>0" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4696
          "r<e1" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4697
          pp_holo:"pp holomorphic_on cball p r" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4698
          pp_po:"(\<forall>w\<in>cball p r-{p}. f w = pp w * (w - p) powr po \<and> pp w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4699
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4700
        have "isolated_singularity_at f p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4701
        proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4702
          have "f holomorphic_on ball p e1 - {p}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4703
            apply (intro holomorphic_on_subset[OF f_holo])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4704
            using e1_avoid \<open>p\<in>pz\<close> unfolding avoid_def pz_def by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4705
          then show ?thesis unfolding isolated_singularity_at_def 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4706
            using \<open>e1>0\<close> analytic_on_open open_delete by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4707
        qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4708
        moreover have "not_essential f p"  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4709
        proof (cases "is_pole f p")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4710
          case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4711
          then show ?thesis unfolding not_essential_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4712
        next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4713
          case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4714
          then have "p\<in>s-poles" using \<open>p\<in>s\<close> poles unfolding pz_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4715
          moreover have "open (s-poles)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4716
            using \<open>open s\<close> 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4717
            apply (elim open_Diff)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4718
            apply (rule finite_imp_closed)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4719
            using finite unfolding pz_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4720
          ultimately have "isCont f p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4721
            using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4722
            by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4723
          then show ?thesis unfolding isCont_def not_essential_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4724
        qed  
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4725
        moreover have "\<exists>\<^sub>F w in at p. f w \<noteq> 0 "
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4726
        proof (rule ccontr)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4727
          assume "\<not> (\<exists>\<^sub>F w in at p. f w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4728
          then have "\<forall>\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4729
          then obtain rr where "rr>0" "\<forall>w\<in>ball p rr - {p}. f w =0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4730
            unfolding eventually_at by (auto simp add:dist_commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4731
          then have "ball p rr - {p} \<subseteq> {w\<in>ball p rr-{p}. f w=0}" by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4732
          moreover have "infinite (ball p rr - {p})" using \<open>rr>0\<close> using finite_imp_not_open by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4733
          ultimately have "infinite {w\<in>ball p rr-{p}. f w=0}" using infinite_super by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4734
          then have "infinite pz"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4735
            unfolding pz_def infinite_super by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4736
          then show False using \<open>finite pz\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4737
        qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4738
        ultimately obtain r where "pp p \<noteq> 0" and r:"r>0" "pp holomorphic_on cball p r" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4739
                  "(\<forall>w\<in>cball p r - {p}. f w = pp w * (w - p) powr of_int po \<and> pp w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4740
          using zorder_exist[of f p,folded po_def pp_def] by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4741
        define r1 where "r1=min r e1 / 2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4742
        have "r1<e1" unfolding r1_def using \<open>e1>0\<close> \<open>r>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4743
        moreover have "r1>0" "pp holomorphic_on cball p r1" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4744
                  "(\<forall>w\<in>cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \<and> pp w \<noteq> 0)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4745
          unfolding r1_def using \<open>e1>0\<close> r by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4746
        ultimately show ?thesis using that \<open>pp p\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4747
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4748
      
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4749
      define e2 where "e2 \<equiv> r/2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4750
      have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4751
      define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4752
      define prin where "prin \<equiv> \<lambda>w. po * h w / (w - p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4753
      have "((\<lambda>w.  prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4754
      proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4755
        have "ball p r \<subseteq> s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4756
          using \<open>r<e1\<close> avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4757
        then have "cball p e2 \<subseteq> s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4758
          using \<open>r>0\<close> unfolding e2_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4759
        then have "(\<lambda>w. po * h w) holomorphic_on cball p e2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4760
          using h_holo by (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4761
        then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4762
          using Cauchy_integral_circlepath_simple[folded c_def, of "\<lambda>w. po * h w"] \<open>e2>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4763
          unfolding prin_def by (auto simp add: mult.assoc)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4764
        have "anal holomorphic_on ball p r" unfolding anal_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4765
          using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close> \<open>pp p\<noteq>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4766
          by (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4767
        then show "(anal has_contour_integral 0) (circlepath p e2)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4768
          using e2_def \<open>r>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4769
          by (auto elim!: Cauchy_theorem_disc_simple)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4770
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4771
      then have "cont ff' p e2" unfolding cont_def po_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4772
      proof (elim has_contour_integral_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4773
        fix w assume "w \<in> path_image (circlepath p e2)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4774
        then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4775
        define wp where "wp \<equiv> w-p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4776
        have "wp\<noteq>0" and "pp w \<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4777
          unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4778
        moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4779
        proof (rule DERIV_imp_deriv)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4780
          have "(pp has_field_derivative (deriv pp w)) (at w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4781
            using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4782
            by (meson open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4783
          then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4784
                  + deriv pp w * (w - p) powr of_int po) (at w)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4785
            unfolding f'_def using \<open>w\<noteq>p\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4786
            apply (auto intro!: derivative_eq_intros(35) DERIV_cong[OF has_field_derivative_powr_of_int])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4787
            by (auto intro: derivative_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4788
        qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4789
        ultimately show "prin w + anal w = ff' w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4790
          unfolding ff'_def prin_def anal_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4791
          apply simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4792
          apply (unfold f'_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4793
          apply (fold wp_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4794
          apply (auto simp add:field_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4795
          by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4796
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4797
      then have "cont ff p e2" unfolding cont_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4798
      proof (elim has_contour_integral_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4799
        fix w assume "w \<in> path_image (circlepath p e2)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4800
        then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4801
        have "deriv f' w =  deriv f w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4802
        proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4803
          show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4804
            by (auto intro!: holomorphic_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4805
        next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4806
          have "ball p e1 - {p} \<subseteq> s - poles"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4807
            using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4808
            by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4809
          then have "ball p r - {p} \<subseteq> s - poles" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4810
            apply (elim dual_order.trans)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4811
            using \<open>r<e1\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4812
          then show "f holomorphic_on ball p r - {p}" using f_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4813
            by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4814
        next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4815
          show "open (ball p r - {p})" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4816
          show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4817
        next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4818
          fix x assume "x \<in> ball p r - {p}"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4819
          then show "f' x = f x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4820
            using pp_po unfolding f'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4821
        qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4822
        moreover have " f' w  =  f w "
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4823
          using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4824
          unfolding f'_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4825
        ultimately show "ff' w = ff w"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4826
          unfolding ff'_def ff_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4827
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4828
      moreover have "cball p e2 \<subseteq> ball p e1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4829
        using \<open>0 < r\<close> \<open>r<e1\<close> e2_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4830
      ultimately show ?thesis using \<open>e2>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4831
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4832
    then obtain e2 where e2:"p\<in>pz \<longrightarrow> e2>0 \<and> cball p e2 \<subseteq> ball p e1 \<and> cont ff p e2"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4833
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4834
    define e4 where "e4 \<equiv> if p\<in>pz then e2 else  e1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4835
    have "e4>0" using e2 \<open>e1>0\<close> unfolding e4_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4836
    moreover have "avoid p e4" using e2 \<open>e1>0\<close> e1_avoid unfolding e4_def avoid_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4837
    moreover have "p\<in>pz \<longrightarrow> cont ff p e4"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4838
      by (auto simp add: e2 e4_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4839
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4840
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4841
  then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4842
      \<and> (p\<in>pz \<longrightarrow> cont ff p (get_e p))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4843
    by metis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4844
  define ci where "ci \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4845
  define w where "w \<equiv> \<lambda>p. winding_number g p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4846
  have "contour_integral g ff = (\<Sum>p\<in>pz. w p * ci p)" unfolding ci_def w_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4847
  proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4848
        path_img homo])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4849
    have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4850
    then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4851
      by (auto intro!: holomorphic_intros simp add:pz_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4852
  next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4853
    show "\<forall>p\<in>s. 0 < get_e p \<and> (\<forall>w\<in>cball p (get_e p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pz))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4854
      using get_e using avoid_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4855
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4856
  also have "... = (\<Sum>p\<in>pz. c * w p * h p * zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4857
  proof (rule sum.cong[of pz pz,simplified])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4858
    fix p assume "p \<in> pz"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4859
    show "w p * ci p = c * w p * h p * (zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4860
    proof (cases "p\<in>s")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4861
      assume "p \<in> s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4862
      have "ci p = c * h p * (zorder f p)" unfolding ci_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4863
        apply (rule contour_integral_unique)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4864
        using get_e \<open>p\<in>s\<close> \<open>p\<in>pz\<close> unfolding cont_def by (metis mult.assoc mult.commute)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4865
      thus ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4866
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4867
      assume "p\<notin>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4868
      then have "w p=0" using homo unfolding w_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4869
      then show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4870
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4871
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4872
  also have "... = c*(\<Sum>p\<in>pz. w p * h p * zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4873
    unfolding sum_distrib_left by (simp add:algebra_simps)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4874
  finally have "contour_integral g ff = c * (\<Sum>p\<in>pz. w p * h p * of_int (zorder f p))" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4875
  then show ?thesis unfolding ff_def c_def w_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4876
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4877
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4878
subsection \<open>Rouche's theorem \<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4879
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4880
theorem Rouche_theorem:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4881
  fixes f g::"complex \<Rightarrow> complex" and s:: "complex set"
68629
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  4882
  defines "fg\<equiv>(\<lambda>p. f p + g p)"
f36858fdf768 Tagged Conformal_Mappings in HOL-Analysis
Wenda Li <wl302@cam.ac.uk>
parents: 68359
diff changeset
  4883
  defines "zeros_fg\<equiv>{p. fg p = 0}" and "zeros_f\<equiv>{p. f p = 0}"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4884
  assumes
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4885
    "open s" and "connected s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4886
    "finite zeros_fg" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4887
    "finite zeros_f" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4888
    f_holo:"f holomorphic_on s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4889
    g_holo:"g holomorphic_on s" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4890
    "valid_path \<gamma>" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4891
    loop:"pathfinish \<gamma> = pathstart \<gamma>" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4892
    path_img:"path_image \<gamma> \<subseteq> s " and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4893
    path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4894
    homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number \<gamma> z = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4895
  shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4896
          = (\<Sum>p\<in>zeros_f. winding_number \<gamma> p * zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4897
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4898
  have path_fg:"path_image \<gamma> \<subseteq> s - zeros_fg"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4899
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4900
    have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4901
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4902
      have "cmod (f z) > cmod (g z)" using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4903
      moreover have "f z = - g z"  using \<open>f z + g z =0\<close> by (simp add: eq_neg_iff_add_eq_0)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4904
      then have "cmod (f z) = cmod (g z)" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4905
      ultimately show False by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4906
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4907
    then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4908
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4909
  have path_f:"path_image \<gamma> \<subseteq> s - zeros_f"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4910
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4911
    have False when "z\<in>path_image \<gamma>" and "f z =0" for z
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4912
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4913
      have "cmod (g z) < cmod (f z) " using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4914
      then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4915
      then show False by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4916
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4917
    then show ?thesis unfolding zeros_f_def using path_img by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4918
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4919
  define w where "w \<equiv> \<lambda>p. winding_number \<gamma> p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4920
  define c where "c \<equiv> 2 * complex_of_real pi * \<i>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4921
  define h where "h \<equiv> \<lambda>p. g p / f p + 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4922
  obtain spikes
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4923
    where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4924
    using \<open>valid_path \<gamma>\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4925
    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4926
  have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4927
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4928
    have outside_img:"0 \<in> outside (path_image (h o \<gamma>))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4929
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4930
      have "h p \<in> ball 1 1" when "p\<in>path_image \<gamma>" for p
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4931
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4932
        have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4933
          apply (cases "cmod (f p) = 0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4934
          by (auto simp add: norm_divide)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4935
        then show ?thesis unfolding h_def by (auto simp add:dist_complex_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4936
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4937
      then have "path_image (h o \<gamma>) \<subseteq> ball 1 1"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4938
        by (simp add: image_subset_iff path_image_compose)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4939
      moreover have " (0::complex) \<notin> ball 1 1" by (simp add: dist_norm)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4940
      ultimately show "?thesis"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4941
        using  convex_in_outside[of "ball 1 1" 0] outside_mono by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4942
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4943
    have valid_h:"valid_path (h \<circ> \<gamma>)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4944
    proof (rule valid_path_compose_holomorphic[OF \<open>valid_path \<gamma>\<close> _ _ path_f])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4945
      show "h holomorphic_on s - zeros_f"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4946
        unfolding h_def using f_holo g_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4947
        by (auto intro!: holomorphic_intros simp add:zeros_f_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4948
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4949
      show "open (s - zeros_f)" using \<open>finite zeros_f\<close> \<open>open s\<close> finite_imp_closed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4950
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4951
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4952
    have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4953
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4954
      have "0 \<notin> path_image (h \<circ> \<gamma>)" using outside_img by (simp add: outside_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4955
      then have "((\<lambda>z. 1/z) has_contour_integral c * winding_number (h \<circ> \<gamma>) 0) (h \<circ> \<gamma>)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4956
        using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4957
        unfolding c_def by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4958
      moreover have "winding_number (h o \<gamma>) 0 = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4959
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4960
        have "0 \<in> outside (path_image (h \<circ> \<gamma>))" using outside_img .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4961
        moreover have "path (h o \<gamma>)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4962
          using valid_h  by (simp add: valid_path_imp_path)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4963
        moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4964
          by (simp add: loop pathfinish_compose pathstart_compose)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4965
        ultimately show ?thesis using winding_number_zero_in_outside by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4966
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4967
      ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4968
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4969
    moreover have "vector_derivative (h \<circ> \<gamma>) (at x) = vector_derivative \<gamma> (at x) * deriv h (\<gamma> x)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4970
      when "x\<in>{0..1} - spikes" for x
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4971
    proof (rule vector_derivative_chain_at_general)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4972
      show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4973
    next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4974
      define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4975
      define t where "t \<equiv> \<gamma> x"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4976
      have "f t\<noteq>0" unfolding zeros_f_def t_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4977
        by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4978
      moreover have "t\<in>s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4979
        using contra_subsetD path_image_def path_fg t_def that by fastforce
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4980
      ultimately have "(h has_field_derivative der t) (at t)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4981
        unfolding h_def der_def using g_holo f_holo \<open>open s\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4982
        by (auto intro!: holomorphic_derivI derivative_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4983
      then show "h field_differentiable at (\<gamma> x)" 
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4984
        unfolding t_def field_differentiable_def by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4985
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4986
    then have " ((/) 1 has_contour_integral 0) (h \<circ> \<gamma>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4987
                  = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4988
      unfolding has_contour_integral
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4989
      apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4990
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4991
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4992
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4993
  then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4994
    using  contour_integral_unique by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4995
  moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4996
      + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4997
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4998
    have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  4999
    proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5000
      show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5001
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5002
      then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5003
        using f_holo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5004
        by (auto intro!: holomorphic_intros simp add:zeros_f_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5005
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5006
    moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5007
      using h_contour
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5008
      by (simp add: has_contour_integral_integrable)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5009
    ultimately have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x + deriv h x / h x) =
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5010
                        contour_integral \<gamma> (\<lambda>p. deriv f p / f p) + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5011
      using contour_integral_add[of "(\<lambda>p. deriv f p / f p)" \<gamma> "(\<lambda>p. deriv h p / h p)" ]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5012
      by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5013
    moreover have "deriv fg p / fg p =  deriv f p / f p + deriv h p / h p"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5014
                      when "p\<in> path_image \<gamma>" for p
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5015
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5016
      have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5017
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5018
      have "h p\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5019
      proof (rule ccontr)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5020
        assume "\<not> h p \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5021
        then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5022
        then have "cmod (g p/f p) = 1" by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5023
        moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5024
          apply (cases "cmod (f p) = 0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5025
          by (auto simp add: norm_divide)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5026
        ultimately show False by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5027
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5028
      have der_fg:"deriv fg p =  deriv f p + deriv g p" unfolding fg_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5029
        using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _  \<open>open s\<close>] path_img that
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5030
        by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5031
      have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5032
      proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5033
        define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5034
        have "p\<in>s" using path_img that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5035
        then have "(h has_field_derivative der p) (at p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5036
          unfolding h_def der_def using g_holo f_holo \<open>open s\<close> \<open>f p\<noteq>0\<close>
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5037
          by (auto intro!: derivative_eq_intros holomorphic_derivI)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5038
        then show ?thesis unfolding der_def using DERIV_imp_deriv by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5039
      qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5040
      show ?thesis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5041
        apply (simp only:der_fg der_h)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5042
        apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5043
        by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5044
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5045
    then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5046
                  = contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5047
      by (elim contour_integral_eq)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5048
    ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5049
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5050
  moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5051
    unfolding c_def zeros_fg_def w_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5052
  proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5053
        , of _ "{}" "\<lambda>_. 1",simplified])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5054
    show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5055
    show "path_image \<gamma> \<subseteq> s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5056
    show " finite {p. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5057
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5058
  moreover have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x) = c * (\<Sum>p\<in>zeros_f. w p * zorder f p)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5059
    unfolding c_def zeros_f_def w_def
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5060
  proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5061
        , of _ "{}" "\<lambda>_. 1",simplified])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5062
    show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5063
    show "path_image \<gamma> \<subseteq> s - {p. f p = 0}" using path_f unfolding zeros_f_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5064
    show " finite {p. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5065
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5066
  ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5067
    by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67443
diff changeset
  5068
  then show ?thesis unfolding c_def using w_def by auto
66456
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  5069
qed
621897f47fab Various lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66394
diff changeset
  5070
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  5071
end