src/HOL/Complex_Analysis/Laurent_Convergence.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82541 5160b68e78a9
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
theory Laurent_Convergence
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
  imports "HOL-Computational_Algebra.Formal_Laurent_Series" "HOL-Library.Landau_Symbols"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
          Residue_Theorem
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
begin
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
definition%important fls_conv_radius :: "complex fls \<Rightarrow> ereal" where
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
  "fls_conv_radius f = fps_conv_radius (fls_regpart f)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
definition%important eval_fls :: "complex fls \<Rightarrow> complex \<Rightarrow> complex" where
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
  "eval_fls F z = eval_fps (fls_base_factor_to_fps F) z * z powi fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
definition\<^marker>\<open>tag important\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
  has_laurent_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex fls \<Rightarrow> bool"
80914
d97fdabd9e2b standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents: 79945
diff changeset
    16
  (infixl \<open>has'_laurent'_expansion\<close> 60)
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
  where "(f has_laurent_expansion F) \<longleftrightarrow>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
            fls_conv_radius F > 0 \<and> eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
lemma has_laurent_expansion_schematicI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
  "f has_laurent_expansion F \<Longrightarrow> F = G \<Longrightarrow> f has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
  by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
lemma has_laurent_expansion_cong:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    25
  assumes "eventually (\<lambda>x. f x = g x) (at 0)" "F = G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  shows   "(f has_laurent_expansion F) \<longleftrightarrow> (g has_laurent_expansion G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
  have "eventually (\<lambda>z. eval_fls F z = g z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
    if "eventually (\<lambda>z. eval_fls F z = f z) (at 0)" "eventually (\<lambda>x. f x = g x) (at 0)" for f g
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
    using that by eventually_elim auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  from this[of f g] this[of g f] show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
    using assms by (auto simp: eq_commute has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
lemma has_laurent_expansion_cong':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  assumes "eventually (\<lambda>x. f x = g x) (at z)" "F = G" "z = z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
  shows   "((\<lambda>x. f (z + x)) has_laurent_expansion F) \<longleftrightarrow> ((\<lambda>x. g (z' + x)) has_laurent_expansion G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
  by (intro has_laurent_expansion_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
     (use assms in \<open>auto simp: at_to_0' eventually_filtermap add_ac\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
lemma fls_conv_radius_altdef:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  "fls_conv_radius F = fps_conv_radius (fls_base_factor_to_fps F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
  have "conv_radius (\<lambda>n. fls_nth F (int n)) = conv_radius (\<lambda>n. fls_nth F (int n + fls_subdegree F))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  proof (cases "fls_subdegree F \<ge> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
    case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
    hence "conv_radius (\<lambda>n. fls_nth F (int n + fls_subdegree F)) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
           conv_radius (\<lambda>n. fls_nth F (int (n + nat (fls_subdegree F))))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
      by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
      by (subst (asm) conv_radius_shift) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
    case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
    hence "conv_radius (\<lambda>n. fls_nth F (int n)) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
           conv_radius (\<lambda>n. fls_nth F (fls_subdegree F + int (n + nat (-fls_subdegree F))))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
      by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
      by (subst (asm) conv_radius_shift) (auto simp: add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
    by (simp add: fls_conv_radius_def fps_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
lemma eval_fps_of_nat [simp]: "eval_fps (of_nat n) z = of_nat n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
  and eval_fps_of_int [simp]: "eval_fps (of_int m) z = of_int m"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  by (simp_all flip: fps_of_nat fps_of_int)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
lemma fps_conv_radius_of_nat [simp]: "fps_conv_radius (of_nat n) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  and fps_conv_radius_of_int [simp]: "fps_conv_radius (of_int m) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  by (simp_all flip: fps_of_nat fps_of_int)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
lemma fps_conv_radius_fls_regpart: "fps_conv_radius (fls_regpart F) = fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
  by (simp add: fls_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
lemma fls_conv_radius_0 [simp]: "fls_conv_radius 0 = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  and fls_conv_radius_1 [simp]: "fls_conv_radius 1 = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
  and fls_conv_radius_const [simp]: "fls_conv_radius (fls_const c) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
  and fls_conv_radius_numeral [simp]: "fls_conv_radius (numeral num) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  and fls_conv_radius_of_nat [simp]: "fls_conv_radius (of_nat n) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  and fls_conv_radius_of_int [simp]: "fls_conv_radius (of_int m) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  and fls_conv_radius_X [simp]: "fls_conv_radius fls_X = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  and fls_conv_radius_X_inv [simp]: "fls_conv_radius fls_X_inv = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
  and fls_conv_radius_X_intpow [simp]: "fls_conv_radius (fls_X_intpow m) = \<infinity>"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
  by (simp_all add: fls_conv_radius_def fls_X_intpow_regpart)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
lemma fls_conv_radius_shift [simp]: "fls_conv_radius (fls_shift n F) = fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  unfolding fls_conv_radius_altdef by (subst fls_base_factor_to_fps_shift) (rule refl)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
lemma fls_conv_radius_fps_to_fls [simp]: "fls_conv_radius (fps_to_fls F) = fps_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  by (simp add: fls_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
lemma fls_conv_radius_deriv [simp]: "fls_conv_radius (fls_deriv F) \<ge> fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  have "fls_conv_radius (fls_deriv F) = fps_conv_radius (fls_regpart (fls_deriv F))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
    by (simp add: fls_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  also have "fls_regpart (fls_deriv F) = fps_deriv (fls_regpart F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
    by (intro fps_ext) (auto simp: add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
  also have "fps_conv_radius \<dots> \<ge> fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
    by (simp add: fls_conv_radius_def fps_conv_radius_deriv)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
lemma fls_conv_radius_uminus [simp]: "fls_conv_radius (-F) = fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  by (simp add: fls_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
lemma fls_conv_radius_add: "fls_conv_radius (F + G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  by (simp add: fls_conv_radius_def fps_conv_radius_add)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
lemma fls_conv_radius_diff: "fls_conv_radius (F - G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  by (simp add: fls_conv_radius_def fps_conv_radius_diff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
lemma fls_conv_radius_mult: "fls_conv_radius (F * G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
proof (cases "F = 0 \<or> G = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
  have "fls_conv_radius (F * G) = fps_conv_radius (fls_regpart (fls_shift (fls_subdegree F + fls_subdegree G) (F * G)))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    by (simp add: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  also have "fls_regpart (fls_shift (fls_subdegree F + fls_subdegree G) (F * G)) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
             fls_base_factor_to_fps F * fls_base_factor_to_fps G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    by (simp add: fls_times_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  also have "fps_conv_radius \<dots> \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
    unfolding fls_conv_radius_altdef by (rule fps_conv_radius_mult)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
lemma fps_conv_radius_add_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F + G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  using fps_conv_radius_add[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
lemma fps_conv_radius_diff_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
  "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F - G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  using fps_conv_radius_diff[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
lemma fps_conv_radius_mult_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F * G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  using fps_conv_radius_mult[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
lemma fls_conv_radius_add_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F + G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  using fls_conv_radius_add[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
lemma fls_conv_radius_diff_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F - G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
  using fls_conv_radius_diff[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
lemma fls_conv_radius_mult_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
  "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F * G) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  using fls_conv_radius_mult[of F G] by (simp add: min_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
lemma fls_conv_radius_power: "fls_conv_radius (F ^ n) \<ge> fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  by (induction n) (auto intro!: fls_conv_radius_mult_ge)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
lemma eval_fls_0 [simp]: "eval_fls 0 z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
  and eval_fls_1 [simp]: "eval_fls 1 z = 1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  and eval_fls_const [simp]: "eval_fls (fls_const c) z = c"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
  and eval_fls_numeral [simp]: "eval_fls (numeral num) z = numeral num"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
  and eval_fls_of_nat [simp]: "eval_fls (of_nat n) z = of_nat n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
  and eval_fls_of_int [simp]: "eval_fls (of_int m) z = of_int m"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
  and eval_fls_X [simp]: "eval_fls fls_X z = z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
  and eval_fls_X_intpow [simp]: "eval_fls (fls_X_intpow m) z = z powi m"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
  by (simp_all add: eval_fls_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
lemma eval_fls_at_0: "eval_fls F 0 = (if fls_subdegree F \<ge> 0 then fls_nth F 0 else 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  by (cases "fls_subdegree F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
     (simp_all add: eval_fls_def fls_regpart_def eval_fps_at_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
lemma eval_fps_to_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
  assumes "norm z < fps_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
  shows   "eval_fls (fps_to_fls F) z = eval_fps F z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
proof (cases "F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  case [simp]: False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  have "eval_fps F z = eval_fps (unit_factor F * normalize F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
    by (metis unit_factor_mult_normalize)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
  also have "\<dots> = eval_fps (unit_factor F * fps_X ^ subdegree F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  also have "\<dots> = eval_fps (unit_factor F) z * z ^ subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
    using assms by (subst eval_fps_mult) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  also have "\<dots> = eval_fls (fps_to_fls F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
    unfolding eval_fls_def fls_base_factor_to_fps_to_fls fls_subdegree_fls_to_fps
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
              power_int_of_nat ..
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
  finally show ?thesis ..
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
lemma eval_fls_shift:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  assumes [simp]: "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
  shows   "eval_fls (fls_shift n F) z = eval_fls F z * z powi -n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
proof (cases "F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
  case [simp]: False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  unfolding eval_fls_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
  by (subst fls_base_factor_to_fps_shift, subst fls_shift_subdegree[OF \<open>F \<noteq> 0\<close>], subst power_int_diff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
     (auto simp: power_int_minus divide_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
lemma eval_fls_add:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G" "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  shows   "eval_fls (F + G) z = eval_fls F z + eval_fls G z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
  using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
proof (induction "fls_subdegree F" "fls_subdegree G" arbitrary: F G rule: linorder_wlog)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
  case (sym F G)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
  show ?case
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
    using sym(1)[of G F] sym(2-) by (simp add: add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  case (le F G)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
  show ?case
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
  proof (cases "F = 0 \<or> G = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
    hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
      by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
    note [simp] = \<open>z \<noteq> 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
    define F' G' where "F' = fls_base_factor_to_fps F" "G' = fls_base_factor_to_fps G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
    define m n where "m = fls_subdegree F" "n = fls_subdegree G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
    have "m \<le> n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
      using le by (auto simp: m_n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
    have conv1: "ereal (cmod z) < fps_conv_radius F'" "ereal (cmod z) < fps_conv_radius G'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
      using assms le by (simp_all add: F'_G'_def fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
    have conv2: "ereal (cmod z) < fps_conv_radius (G' * fps_X ^ nat (n - m))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
      using conv1 by (intro less_le_trans[OF _ fps_conv_radius_mult]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
    have conv3: "ereal (cmod z) < fps_conv_radius (F' + G' * fps_X ^ nat (n - m))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
      using conv1 conv2 by (intro less_le_trans[OF _ fps_conv_radius_add]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
    have "eval_fls F z + eval_fls G z = eval_fps F' z * z powi m + eval_fps G' z * z powi n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
      unfolding eval_fls_def m_n_def[symmetric] F'_G'_def[symmetric]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
      by (simp add: power_int_add algebra_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
    also have "\<dots> = (eval_fps F' z + eval_fps G' z * z powi (n - m)) * z powi m"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
      by (simp add: algebra_simps power_int_diff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
    also have "eval_fps G' z * z powi (n - m) = eval_fps (G' * fps_X ^ nat (n - m)) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
      using assms \<open>m \<le> n\<close> conv1 by (subst eval_fps_mult) (auto simp: power_int_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
    also have "eval_fps F' z + \<dots> = eval_fps (F' + G' * fps_X ^ nat (n - m)) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
      using conv1 conv2 by (subst eval_fps_add) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
    also have "\<dots> = eval_fls (fps_to_fls (F' + G' * fps_X ^ nat (n - m))) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
      using conv3 by (subst eval_fps_to_fls) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
    also have "\<dots> * z powi m = eval_fls (fls_shift (-m) (fps_to_fls (F' + G' * fps_X ^ nat (n - m)))) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
      by (subst eval_fls_shift) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
    also have "fls_shift (-m) (fps_to_fls (F' + G' * fps_X ^ nat (n - m))) = F + G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
      using \<open>m \<le> n\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
      by (simp add: fls_times_fps_to_fls fps_to_fls_power fls_X_power_conv_shift_1
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
                    fls_shifted_times_simps F'_G'_def m_n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    finally show ?thesis ..
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
  qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
lemma eval_fls_minus:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
  assumes "ereal (norm z) < fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
  shows   "eval_fls (-F) z = -eval_fls F z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
  using assms by (simp add: eval_fls_def eval_fps_minus fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
lemma eval_fls_diff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
  assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
     and [simp]: "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  shows   "eval_fls (F - G) z = eval_fls F z - eval_fls G z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   253
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   254
  have "eval_fls (F + (-G)) z = eval_fls F z - eval_fls G z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
    using assms by (subst eval_fls_add) (auto simp: eval_fls_minus)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
lemma eval_fls_mult:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
  assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G" "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  shows   "eval_fls (F * G) z = eval_fls F z * eval_fls G z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
proof (cases "F = 0 \<or> G = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
  note [simp] = \<open>z \<noteq> 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  define F' G' where "F' = fls_base_factor_to_fps F" "G' = fls_base_factor_to_fps G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
  define m n where "m = fls_subdegree F" "n = fls_subdegree G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
  have "eval_fls F z * eval_fls G z = (eval_fps F' z * eval_fps G' z) * z powi (m + n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    unfolding eval_fls_def m_n_def[symmetric] F'_G'_def[symmetric]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
    by (simp add: power_int_add algebra_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
  also have "\<dots> = eval_fps (F' * G') z * z powi (m + n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
    using assms by (subst eval_fps_mult) (auto simp: F'_G'_def fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
  also have "\<dots> = eval_fls (F * G) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    by (simp add: eval_fls_def F'_G'_def m_n_def) (simp add: fls_times_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
  finally show ?thesis ..
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
lemma eval_fls_power:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
  assumes "ereal (norm z) < fls_conv_radius F" "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
  shows   "eval_fls (F ^ n) z = eval_fls F z ^ n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
proof (induction n)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
  case (Suc n)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
  have "eval_fls (F ^ Suc n) z = eval_fls (F * F ^ n) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
  also have "\<dots> = eval_fls F z * eval_fls (F ^ n) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
    using assms by (subst eval_fls_mult) (auto intro!: less_le_trans[OF _ fls_conv_radius_power])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
  finally show ?case
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    using Suc by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
82338
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   293
lemma eval_fls_eq:
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   294
  assumes "N \<le> fls_subdegree F" "fls_subdegree F \<ge> 0 \<or> z \<noteq> 0"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   295
  assumes "(\<lambda>n. fls_nth F (int n + N) * z powi (int n + N)) sums S"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   296
  shows   "eval_fls F z = S"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   297
proof (cases "z = 0")
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   298
  case [simp]: True
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   299
  have "(\<lambda>n. fls_nth F (int n + N) * z powi (int n + N)) =
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   300
        (\<lambda>n. if n \<in> (if N \<le> 0 then {nat (-N)} else {}) then fls_nth F (int n + N) else 0)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   301
    by (auto simp: fun_eq_iff split: if_splits)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   302
  also have "\<dots> sums (\<Sum>n\<in>(if N \<le> 0 then {nat (-N)} else {}). fls_nth F (int n + N))"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   303
    by (rule sums_If_finite_set) auto
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   304
  also have "\<dots> = fls_nth F 0"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   305
    using assms by auto
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   306
  also have "\<dots> = eval_fls F z"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   307
    using assms by (auto simp: eval_fls_def eval_fps_at_0 power_int_0_left_if)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   308
  finally show ?thesis 
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   309
    using assms by (simp add: sums_iff)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   310
next
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   311
  case [simp]: False
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   312
  define N' where "N' = fls_subdegree F"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   313
  define d where "d = nat (N' - N)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   314
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   315
  have "(\<lambda>n. fls_nth F (int n + N) * z powi (int n + N)) sums S"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   316
    by fact
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   317
  also have "?this \<longleftrightarrow> (\<lambda>n. fls_nth F (int (n+d) + N) * z powi (int (n+d) + N)) sums S"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   318
    by (rule sums_zero_iff_shift [symmetric]) (use assms in \<open>auto simp: d_def N'_def\<close>)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   319
  also have "(\<lambda>n. int (n+d) + N) = (\<lambda>n. int n + N')"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   320
    using assms by (auto simp: N'_def d_def)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   321
  finally have "(\<lambda>n. fls_nth F (int n + N') * z powi (int n + N')) sums S" .
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   322
  hence "(\<lambda>n. z powi (-N') * (fls_nth F (int n + N') * z powi (int n + N'))) sums (z powi (-N') * S)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   323
    by (intro sums_mult)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   324
  hence "(\<lambda>n. fls_nth F (int n + N') * z ^ n) sums (z powi (-N') * S)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   325
    by (simp add: power_int_add power_int_minus field_simps)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   326
  thus ?thesis
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   327
    by (simp add: eval_fls_def eval_fps_def sums_iff power_int_minus N'_def)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   328
qed
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80948
diff changeset
   329
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
lemma norm_summable_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
  "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. norm (fls_nth f n * z ^ n))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
  using norm_summable_fps[of z "fls_regpart f"] by (simp add: fls_conv_radius_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
lemma norm_summable_fls':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
  "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. norm (fls_nth f (n + fls_subdegree f) * z ^ n))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
  using norm_summable_fps[of z "fls_base_factor_to_fps f"] by (simp add: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
lemma summable_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   339
  "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. fls_nth f n * z ^ n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
  by (rule summable_norm_cancel[OF norm_summable_fls])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
theorem sums_eval_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
  fixes f
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
  defines "n \<equiv> fls_subdegree f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
  assumes "norm z < fls_conv_radius f" and "z \<noteq> 0 \<or> n \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
  shows   "(\<lambda>k. fls_nth f (int k + n) * z powi (int k + n)) sums eval_fls f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
proof (cases "z = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
  case [simp]: False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
  have "(\<lambda>k. fps_nth (fls_base_factor_to_fps f) k * z ^ k * z powi n) sums
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
          (eval_fps (fls_base_factor_to_fps f) z * z powi n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
    using assms(2) by (intro sums_eval_fps sums_mult2) (auto simp: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
    by (simp add: power_int_add n_def eval_fls_def mult_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  case [simp]: True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
  with assms have "n \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
  have "(\<lambda>k. fls_nth f (int k + n) * z powi (int k + n)) sums
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
          (\<Sum>k\<in>(if n \<le> 0 then {nat (-n)} else {}). fls_nth f (int k + n) * z powi (int k + n))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
    by (intro sums_finite) (auto split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
  also have "\<dots> = eval_fls f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
    using \<open>n \<ge> 0\<close> by (auto simp: eval_fls_at_0 n_def not_le)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
lemma holomorphic_on_eval_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
  fixes f
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
  defines "n \<equiv> fls_subdegree f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
  assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
  shows   "eval_fls f holomorphic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
proof (cases "n \<ge> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
  have "eval_fls f = (\<lambda>z. eval_fps (fls_base_factor_to_fps f) z * z ^ nat n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
    using True by (simp add: fun_eq_iff eval_fls_def power_int_def n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
  moreover have "\<dots> holomorphic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
    using True assms(2) by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
  ultimately show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
  show ?thesis using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
    unfolding eval_fls_def by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
lemma holomorphic_on_eval_fls' [holomorphic_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
  assumes "g holomorphic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
  assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
  shows   "(\<lambda>x. eval_fls f (g x)) holomorphic_on A"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   389
  by (meson assms holomorphic_on_compose holomorphic_on_eval_fls holomorphic_transform o_def)
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
lemma continuous_on_eval_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
  fixes f
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
  defines "n \<equiv> fls_subdegree f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
  shows   "continuous_on A (eval_fls f)"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   396
  using assms holomorphic_on_eval_fls holomorphic_on_imp_continuous_on by blast
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
lemma continuous_on_eval_fls' [continuous_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
  fixes f
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
  defines "n \<equiv> fls_subdegree f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
  assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
  assumes "continuous_on A g"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
  shows   "continuous_on A (\<lambda>x. eval_fls f (g x))"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   404
  by (metis assms continuous_on_compose2 continuous_on_eval_fls order.refl)
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
lemmas has_field_derivative_eval_fps' [derivative_intros] =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
  DERIV_chain2[OF has_field_derivative_eval_fps]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
(* TODO: generalise for nonneg subdegree *)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
lemma has_field_derivative_eval_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
  assumes "z \<in> eball 0 (fls_conv_radius f) - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
  shows   "(eval_fls f has_field_derivative eval_fls (fls_deriv f) z) (at z within A)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
  define g where "g = fls_base_factor_to_fps f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
  define n where "n = fls_subdegree f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
  have [simp]: "fps_conv_radius g = fls_conv_radius f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
    by (simp add: fls_conv_radius_altdef g_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
  have conv1: "fps_conv_radius (fps_deriv g * fps_X) \<ge> fls_conv_radius f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
    by (intro fps_conv_radius_mult_ge order.trans[OF _ fps_conv_radius_deriv]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
  have conv2: "fps_conv_radius (of_int n * g) \<ge> fls_conv_radius f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
    by (intro fps_conv_radius_mult_ge) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
  have conv3: "fps_conv_radius (fps_deriv g * fps_X + of_int n * g) \<ge> fls_conv_radius f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
    by (intro fps_conv_radius_add_ge conv1 conv2)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
  have [simp]: "fps_conv_radius g = fls_conv_radius f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
    by (simp add: g_def fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
  have "((\<lambda>z. eval_fps g z * z powi fls_subdegree f) has_field_derivative
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
          (eval_fps (fps_deriv g) z * z powi n + of_int n * z powi (n - 1) * eval_fps g z))
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
          (at z within A)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
    using assms by (auto intro!: derivative_eq_intros simp: n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
  also have "(\<lambda>z. eval_fps g z * z powi fls_subdegree f) = eval_fls f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
    by (simp add: eval_fls_def g_def fun_eq_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
  also have "eval_fps (fps_deriv g) z * z powi n + of_int n * z powi (n - 1) * eval_fps g z =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
             (z * eval_fps (fps_deriv g) z + of_int n * eval_fps g z) * z powi (n - 1)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
    using assms by (auto simp: power_int_diff field_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
  also have "(z * eval_fps (fps_deriv g) z + of_int n * eval_fps g z) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
             eval_fps (fps_deriv g * fps_X + of_int n * g) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
    using conv1 conv2 assms fps_conv_radius_deriv[of g]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
    by (subst eval_fps_add) (auto simp: eval_fps_mult)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
  also have "\<dots> = eval_fls (fps_to_fls (fps_deriv g * fps_X + of_int n * g)) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
    using conv3 assms by (subst eval_fps_to_fls) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
  also have "\<dots> * z powi (n - 1) = eval_fls (fls_shift (1 - n) (fps_to_fls (fps_deriv g * fps_X + of_int n * g))) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
    using assms by (subst eval_fls_shift) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
  also have "fls_shift (1 - n) (fps_to_fls (fps_deriv g * fps_X + of_int n * g)) = fls_deriv f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
    by (intro fls_eqI) (auto simp: g_def n_def algebra_simps eq_commute[of _ "fls_subdegree f"])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
lemma eval_fls_deriv:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  assumes "z \<in> eball 0 (fls_conv_radius F) - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
  shows   "eval_fls (fls_deriv F) z = deriv (eval_fls F) z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   452
  by (metis DERIV_imp_deriv assms has_field_derivative_eval_fls)
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
lemma analytic_on_eval_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
  assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
  shows   "eval_fls f analytic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
proof (rule analytic_on_subset [OF _ assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
  show "eval_fls f analytic_on eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
    using holomorphic_on_eval_fls[OF order.refl]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
    by (subst analytic_on_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
lemma analytic_on_eval_fls' [analytic_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  assumes "g analytic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
  assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  shows   "(\<lambda>x. eval_fls f (g x)) analytic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
  have "eval_fls f \<circ> g analytic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
    by (intro analytic_on_compose[OF assms(1) analytic_on_eval_fls]) (use assms in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
    by (simp add: o_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
lemma continuous_eval_fls [continuous_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
  assumes "z \<in> eball 0 (fls_conv_radius F) - (if fls_subdegree F \<ge> 0 then {} else {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
  shows   "continuous (at z within A) (eval_fls F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  have "isCont (eval_fls F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
    using continuous_on_eval_fls[OF order.refl] assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
    by (subst (asm) continuous_on_eq_continuous_at) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
    using continuous_at_imp_continuous_at_within by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
named_theorems laurent_expansion_intros
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
lemma has_laurent_expansion_imp_asymp_equiv_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
  assumes F: "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
  defines "n \<equiv> fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  shows   "f \<sim>[at 0] (\<lambda>z. fls_nth F n * z powi n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
proof (cases "F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
  thus ?thesis using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   497
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
  case [simp]: False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
  define G where "G = fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
  have "fls_conv_radius F > 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
    using F by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
  hence "isCont (eval_fps G) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
    by (intro continuous_intros) (auto simp: G_def fps_conv_radius_fls_regpart zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
  hence lim: "eval_fps G \<midarrow>0\<rightarrow> eval_fps G 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
    by (meson isContD)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
  have [simp]: "fps_nth G 0 \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
    by (auto simp: G_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
  have "f \<sim>[at 0] eval_fls F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
    using F by (intro asymp_equiv_refl_ev) (auto simp: has_laurent_expansion_def eq_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
  also have "\<dots> = (\<lambda>z. eval_fps G z * z powi n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
    by (intro ext) (simp_all add: eval_fls_def G_def n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
  also have "\<dots> \<sim>[at 0] (\<lambda>z. fps_nth G 0 * z powi n)" using lim
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
    by (intro asymp_equiv_intros tendsto_imp_asymp_equiv_const) (auto simp: eval_fps_at_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
  also have "fps_nth G 0 = fls_nth F n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
    by (simp add: G_def n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
  finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
lemma has_laurent_expansion_imp_asymp_equiv:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
  assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
  defines "n \<equiv> fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
  shows   "f \<sim>[at z] (\<lambda>w. fls_nth F n * (w - z) powi n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
  using has_laurent_expansion_imp_asymp_equiv_0[OF assms(1)] unfolding n_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
  by (simp add: at_to_0[of z] asymp_equiv_filtermap_iff add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
lemmas [tendsto_intros del] = tendsto_power_int
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
lemma has_laurent_expansion_imp_tendsto_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
  assumes F: "f has_laurent_expansion F" and "fls_subdegree F \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  shows   "f \<midarrow>0\<rightarrow> fls_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
proof (rule asymp_equiv_tendsto_transfer)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
  show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<sim>[at 0] f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    by (rule asymp_equiv_symI, rule has_laurent_expansion_imp_asymp_equiv_0) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
  show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<midarrow>0\<rightarrow> fls_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
    by (rule tendsto_eq_intros refl | use assms(2) in simp)+
80948
572970d15ab0 To tiny but maybe useful lemmas (moved in from the AFP, Word_Lib)
paulson <lp15@cam.ac.uk>
parents: 80914
diff changeset
   540
       (use assms(2) in \<open>auto simp: power_int_0_left_if\<close>)
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
lemma has_laurent_expansion_imp_tendsto:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
  assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" and "fls_subdegree F \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
  shows   "f \<midarrow>z\<rightarrow> fls_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
  using has_laurent_expansion_imp_tendsto_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  by (simp add: at_to_0[of z] filterlim_filtermap add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
lemma has_laurent_expansion_imp_filterlim_infinity_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
  assumes F: "f has_laurent_expansion F" and "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
  shows   "filterlim f at_infinity (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
proof (rule asymp_equiv_at_infinity_transfer)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
  have [simp]: "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   554
    using assms(2) by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
  show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<sim>[at 0] f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   556
    by (rule asymp_equiv_symI, rule has_laurent_expansion_imp_asymp_equiv_0) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   557
  show "filterlim (\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) at_infinity (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
    by (rule tendsto_mult_filterlim_at_infinity tendsto_const
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
             filterlim_power_int_neg_at_infinity | use assms(2) in simp)+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
       (auto simp: eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   563
lemma has_laurent_expansion_imp_neg_fls_subdegree:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   564
  assumes F: "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   565
    and infy:"filterlim f at_infinity (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   566
  shows   "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   567
proof (rule ccontr)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   568
  assume asm:"\<not> fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   569
  define ff where "ff=(\<lambda>z. fls_nth F (fls_subdegree F)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   570
                              * z powi fls_subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   571
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   572
  have "(ff \<longlongrightarrow> (if fls_subdegree F =0 then fls_nth F 0 else 0)) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   573
    using asm unfolding ff_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   574
    by (auto intro!: tendsto_eq_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   575
  moreover have "filterlim ff at_infinity (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   576
  proof (rule asymp_equiv_at_infinity_transfer)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   577
    show "f \<sim>[at 0] ff" unfolding ff_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
      using has_laurent_expansion_imp_asymp_equiv_0[OF F] unfolding ff_def .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   579
    show "filterlim f at_infinity (at 0)" by fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   580
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   581
  ultimately show False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   582
    using not_tendsto_and_filterlim_at_infinity[of "at (0::complex)"] by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
lemma has_laurent_expansion_imp_filterlim_infinity:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
  assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" and "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
  shows   "filterlim f at_infinity (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
  using has_laurent_expansion_imp_filterlim_infinity_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
  by (simp add: at_to_0[of z] filterlim_filtermap add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
lemma has_laurent_expansion_imp_is_pole_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
  assumes F: "f has_laurent_expansion F" and "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
  shows   "is_pole f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
  using has_laurent_expansion_imp_filterlim_infinity_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
  by (simp add: is_pole_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
lemma is_pole_0_imp_neg_fls_subdegree:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
  assumes F: "f has_laurent_expansion F" and "is_pole f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
  shows   "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
  using F assms(2) has_laurent_expansion_imp_neg_fls_subdegree is_pole_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
  by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
lemma has_laurent_expansion_imp_is_pole:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
  assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" and "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
  shows   "is_pole f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
  using has_laurent_expansion_imp_is_pole_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
  by (simp add: is_pole_shift_0')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
lemma is_pole_imp_neg_fls_subdegree:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   610
  assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" and "is_pole f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   611
  shows   "fls_subdegree F < 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   612
proof -
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   613
  have "is_pole (\<lambda>x. f (z + x)) 0"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   614
    using assms(2) is_pole_shift_0 by blast
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   615
  then show ?thesis
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   616
    using F is_pole_0_imp_neg_fls_subdegree by blast
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   617
qed
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   618
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
lemma is_pole_fls_subdegree_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
  assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
  shows "is_pole f z \<longleftrightarrow> fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
  using assms is_pole_imp_neg_fls_subdegree has_laurent_expansion_imp_is_pole
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
  by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
lemma
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   627
  shows   has_laurent_expansion_isolated_0: "isolated_singularity_at f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   628
    and   has_laurent_expansion_not_essential_0: "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
  from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   632
  then obtain r where r: "r > 0" "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> eval_fls F z = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
    by (auto simp: eventually_at_filter ball_def eventually_nhds_metric)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   634
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
  have "fls_conv_radius F > 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
    using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   637
  then obtain R :: real where R: "R > 0" "R \<le> min r (fls_conv_radius F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
    using \<open>r > 0\<close> by (metis dual_order.strict_implies_order ereal_dense2 ereal_less(2) min_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
  have "eval_fls F holomorphic_on ball 0 R - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
    using r R by (intro holomorphic_intros ball_eball_mono Diff_mono)  (auto simp: ereal_le_less)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
  also have "?this \<longleftrightarrow> f holomorphic_on ball 0 R - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   643
    using r R by (intro holomorphic_cong) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
  also have "\<dots> \<longleftrightarrow> f analytic_on ball 0 R - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
    by (subst analytic_on_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
  finally show "isolated_singularity_at f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
    unfolding isolated_singularity_at_def using \<open>R > 0\<close> by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   648
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
  show "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
  proof (cases "fls_subdegree F \<ge> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
    case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
    hence "f \<midarrow>0\<rightarrow> fls_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
      by (intro has_laurent_expansion_imp_tendsto_0[OF assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   654
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
      by (auto simp: not_essential_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
    case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
    hence "is_pole f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
      by (intro has_laurent_expansion_imp_is_pole_0[OF assms]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
      by (auto simp: not_essential_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
lemma
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
  assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
  shows   has_laurent_expansion_isolated: "isolated_singularity_at f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
    and   has_laurent_expansion_not_essential: "not_essential f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
  using has_laurent_expansion_isolated_0[OF assms] has_laurent_expansion_not_essential_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
  by (simp_all add: isolated_singularity_at_shift_0 not_essential_shift_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   672
lemma has_laurent_expansion_fps:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
  shows   "f has_laurent_expansion fps_to_fls F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
  from assms have radius: "0 < fps_conv_radius F" and eval: "\<forall>\<^sub>F z in nhds 0. eval_fps F z = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
    by (auto simp: has_fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
  from eval have eval': "\<forall>\<^sub>F z in at 0. eval_fps F z = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   679
    using eventually_at_filter eventually_mono by fastforce
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   680
  moreover have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   681
    using radius by (intro eventually_at_in_open) (auto simp: zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   682
  ultimately have "eventually (\<lambda>z. eval_fls (fps_to_fls F) z = f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   683
    by eventually_elim (auto simp: eval_fps_to_fls)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   684
  thus ?thesis using radius
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   685
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   687
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
lemma has_laurent_expansion_const [simp, intro, laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
  "(\<lambda>_. c) has_laurent_expansion fls_const c"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
lemma has_laurent_expansion_0 [simp, intro, laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
  "(\<lambda>_. 0) has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
lemma has_fps_expansion_0_iff: "f has_fps_expansion 0 \<longleftrightarrow> eventually (\<lambda>z. f z = 0) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
  by (auto simp: has_fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
lemma has_laurent_expansion_1 [simp, intro, laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   700
  "(\<lambda>_. 1) has_laurent_expansion 1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
lemma has_laurent_expansion_numeral [simp, intro, laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
  "(\<lambda>_. numeral n) has_laurent_expansion numeral n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   707
lemma has_laurent_expansion_fps_X_power [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
  "(\<lambda>x. x ^ n) has_laurent_expansion (fls_X_intpow n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   709
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   710
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
lemma has_laurent_expansion_fps_X_power_int [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
  "(\<lambda>x. x powi n) has_laurent_expansion (fls_X_intpow n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   714
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
lemma has_laurent_expansion_fps_X [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
  "(\<lambda>x. x) has_laurent_expansion fls_X"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
  by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
lemma has_laurent_expansion_cmult_left [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   721
  shows   "(\<lambda>x. c * f x) has_laurent_expansion fls_const c * F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   722
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   723
  from assms have "eventually (\<lambda>z. z \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
    by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   725
  moreover from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
  ultimately have "eventually (\<lambda>z. eval_fls (fls_const c * F) z = c * f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
    by eventually_elim (simp_all add: eval_fls_mult)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
  with assms show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
    by (auto simp: has_laurent_expansion_def intro!: less_le_trans[OF _ fls_conv_radius_mult])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   732
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   733
lemma has_laurent_expansion_cmult_right [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
  shows   "(\<lambda>x. f x * c) has_laurent_expansion F * fls_const c"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   737
  have "F * fls_const c = fls_const c * F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
    by (intro fls_eqI) (auto simp: mult.commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
  with has_laurent_expansion_cmult_left [OF assms] show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
    by (simp add: mult.commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
lemma has_fps_expansion_scaleR [fps_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
  fixes F :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
  shows "f has_fps_expansion F \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) has_fps_expansion fps_const (of_real c) * F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
  unfolding scaleR_conv_of_real by (intro fps_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
lemma has_laurent_expansion_scaleR [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   749
  "f has_laurent_expansion F \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) has_laurent_expansion fls_const (of_real c) * F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
  unfolding scaleR_conv_of_real by (intro laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
lemma has_laurent_expansion_minus [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
  shows   "(\<lambda>x. - f x) has_laurent_expansion -F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
  from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
    by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
  moreover from assms have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
  ultimately have "eventually (\<lambda>x. eval_fls (-F) x = -f x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
    by eventually_elim (auto simp: eval_fls_minus)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
  thus ?thesis using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
lemma has_laurent_expansion_add [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
  assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
  shows   "(\<lambda>x. f x + g x) has_laurent_expansion F + G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
  from assms have "0 < min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
  also have "\<dots> \<le> fls_conv_radius (F + G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
    by (rule fls_conv_radius_add)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
  finally have radius: "\<dots> > 0" .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
  from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
                  "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius G) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
    by (intro eventually_at_in_open; force simp: has_laurent_expansion_def zero_ereal_def)+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
  moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
            and "eventually (\<lambda>x. eval_fls G x = g x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
    using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
  ultimately have "eventually (\<lambda>x. eval_fls (F + G) x = f x + g x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
    by eventually_elim (auto simp: eval_fls_add)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
  with radius show ?thesis by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
lemma has_laurent_expansion_diff [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
  shows   "(\<lambda>x. f x - g x) has_laurent_expansion F - G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   789
  using has_laurent_expansion_add[of f F "\<lambda>x. - g x" "-G"] assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
  by (simp add: has_laurent_expansion_minus)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
lemma has_laurent_expansion_mult [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
  assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   794
  shows   "(\<lambda>x. f x * g x) has_laurent_expansion F * G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   795
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
  from assms have "0 < min (fls_conv_radius F) (fls_conv_radius G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
  also have "\<dots> \<le> fls_conv_radius (F * G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
    by (rule fls_conv_radius_mult)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
  finally have radius: "\<dots> > 0" .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   802
  from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   803
                  "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius G) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
    by (intro eventually_at_in_open; force simp: has_laurent_expansion_def zero_ereal_def)+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
  moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
            and "eventually (\<lambda>x. eval_fls G x = g x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
    using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
  ultimately have "eventually (\<lambda>x. eval_fls (F * G) x = f x * g x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
    by eventually_elim (auto simp: eval_fls_mult)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   810
  with radius show ?thesis by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
lemma has_fps_expansion_power [fps_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
  fixes F :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
  shows "f has_fps_expansion F \<Longrightarrow> (\<lambda>x. f x ^ m) has_fps_expansion F ^ m"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
  by (induction m) (auto intro!: fps_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
lemma has_laurent_expansion_power [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
  shows   "(\<lambda>x. f x ^ n) has_laurent_expansion F ^ n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
  by (induction n) (auto intro!: laurent_expansion_intros assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
lemma has_laurent_expansion_sum [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
  assumes "\<And>x. x \<in> I \<Longrightarrow> f x has_laurent_expansion F x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
  shows   "(\<lambda>y. \<Sum>x\<in>I. f x y) has_laurent_expansion (\<Sum>x\<in>I. F x)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
  using assms by (induction I rule: infinite_finite_induct) (auto intro!: laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
lemma has_laurent_expansion_prod [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
  assumes "\<And>x. x \<in> I \<Longrightarrow> f x has_laurent_expansion F x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
  shows   "(\<lambda>y. \<Prod>x\<in>I. f x y) has_laurent_expansion (\<Prod>x\<in>I. F x)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
  using assms by (induction I rule: infinite_finite_induct) (auto intro!: laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   833
lemma has_laurent_expansion_deriv [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   834
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   835
  shows   "deriv f has_laurent_expansion fls_deriv F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  have "eventually (\<lambda>z. z \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
    using assms by (intro eventually_at_in_open)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
                   (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
  moreover from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   842
  then obtain s where "open s" "0 \<in> s" and s: "\<And>w. w \<in> s - {0} \<Longrightarrow> eval_fls F w = f w"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
    by (auto simp: eventually_nhds eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
  hence "eventually (\<lambda>w. w \<in> s - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
    by (intro eventually_at_in_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
  ultimately have "eventually (\<lambda>z. eval_fls (fls_deriv F) z = deriv f z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
  proof eventually_elim
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
    case (elim z)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
    hence "eval_fls (fls_deriv F) z = deriv (eval_fls F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
      by (simp add: eval_fls_deriv)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
    also have "eventually (\<lambda>w. w \<in> s - {0}) (nhds z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
      using elim and \<open>open s\<close> by (intro eventually_nhds_in_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
    hence "eventually (\<lambda>w. eval_fls F w = f w) (nhds z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
      by eventually_elim (use s in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
    hence "deriv (eval_fls F) z = deriv f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
      by (intro deriv_cong_ev refl)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
    finally show ?case .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
  with assms show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
    by (auto simp: has_laurent_expansion_def intro!: less_le_trans[OF _ fls_conv_radius_deriv])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   862
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
lemma has_laurent_expansion_shift [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
  shows   "(\<lambda>x. f x * x powi n) has_laurent_expansion (fls_shift (-n) F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
  have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   868
    using assms by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
  moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
    using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   871
  ultimately have "eventually (\<lambda>x. eval_fls (fls_shift (-n) F) x = f x * x powi n) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   872
    by eventually_elim (auto simp: eval_fls_shift assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   873
  with assms show ?thesis by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   874
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   876
lemma has_laurent_expansion_shift' [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   877
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   878
  shows   "(\<lambda>x. f x * x powi (-n)) has_laurent_expansion (fls_shift n F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
  using has_laurent_expansion_shift[OF assms, of "-n"] by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
lemma has_laurent_expansion_deriv':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
  assumes "open A" "0 \<in> A" "\<And>x. x \<in> A - {0} \<Longrightarrow> (f has_field_derivative f' x) (at x)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
  shows   "f' has_laurent_expansion fls_deriv F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
  have "deriv f has_laurent_expansion fls_deriv F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
    by (intro laurent_expansion_intros assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
  also have "?this \<longleftrightarrow> ?thesis"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
  proof (intro has_laurent_expansion_cong refl)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   891
    have "eventually (\<lambda>z. z \<in> A - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
      by (intro eventually_at_in_open assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   893
    thus "eventually (\<lambda>z. deriv f z = f' z) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   894
      by eventually_elim (auto intro!: DERIV_imp_deriv assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   895
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   896
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   899
definition laurent_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex fls" where
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   900
  "laurent_expansion f z =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
     (if eventually (\<lambda>z. f z = 0) (at z) then 0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   902
      else fls_shift (-zorder f z) (fps_to_fls (fps_expansion (zor_poly f z) z)))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
lemma laurent_expansion_cong:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
  assumes "eventually (\<lambda>w. f w = g w) (at z)" "z = z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   906
  shows   "laurent_expansion f z = laurent_expansion g z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
  unfolding laurent_expansion_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   908
  using zor_poly_cong[OF assms(1,2)] zorder_cong[OF assms] assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
  by (intro if_cong refl) (auto elim: eventually_elim2)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
theorem not_essential_has_laurent_expansion_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
  assumes "isolated_singularity_at f 0" "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
  shows   "f has_laurent_expansion laurent_expansion f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
proof (cases "\<exists>\<^sub>F w in at 0. f w \<noteq> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
  have "(\<lambda>_. 0) has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
  also have "?this \<longleftrightarrow> f has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
    using False by (intro has_laurent_expansion_cong) (auto simp: frequently_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
  finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
    using False by (simp add: laurent_expansion_def frequently_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
  define n where "n = zorder f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
  obtain r where r: "zor_poly f 0 0 \<noteq> 0" "zor_poly f 0 holomorphic_on cball 0 r" "r > 0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   926
                    "\<forall>w\<in>cball 0 r - {0}. f w = zor_poly f 0 w * w powi n \<and>
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
                                         zor_poly f 0 w \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
    using zorder_exist[OF assms True] unfolding n_def by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
  have holo: "zor_poly f 0 holomorphic_on ball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
    by (rule holomorphic_on_subset[OF r(2)]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
  define F where "F = fps_expansion (zor_poly f 0) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
  have F: "zor_poly f 0 has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
    unfolding F_def by (rule has_fps_expansion_fps_expansion[OF _ _ holo]) (use \<open>r > 0\<close> in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
  have "(\<lambda>z. zor_poly f 0 z * z powi n) has_laurent_expansion fls_shift (-n) (fps_to_fls F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
    by (intro laurent_expansion_intros has_laurent_expansion_fps[OF F])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
  also have "?this \<longleftrightarrow> f has_laurent_expansion fls_shift (-n) (fps_to_fls F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
    by (intro has_laurent_expansion_cong refl eventually_mono[OF eventually_at_in_open[of "ball 0 r"]])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
       (use r in \<open>auto simp: complex_powr_of_int\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
  finally show ?thesis using True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
    by (simp add: laurent_expansion_def F_def n_def frequently_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
lemma not_essential_has_laurent_expansion:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
  assumes "isolated_singularity_at f z" "not_essential f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
  shows   "(\<lambda>x. f (z + x)) has_laurent_expansion laurent_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
  from assms(1) have iso:"isolated_singularity_at (\<lambda>x. f (z + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
    by (simp add: isolated_singularity_at_shift_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
  moreover from assms(2) have ness:"not_essential (\<lambda>x. f (z + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
    by (simp add: not_essential_shift_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
  ultimately have "(\<lambda>x. f (z + x)) has_laurent_expansion laurent_expansion (\<lambda>x. f (z + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
    by (rule not_essential_has_laurent_expansion_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
  also have "\<dots> = laurent_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
  proof (cases "\<exists>\<^sub>F w in at z. f w \<noteq> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
    case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
    then have "\<forall>\<^sub>F w in at z. f w = 0" using not_frequently by force
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
    then have "laurent_expansion (\<lambda>x. f (z + x)) 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
      by (smt (verit, best) add.commute eventually_at_to_0 eventually_mono
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
          laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
    moreover have "laurent_expansion f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
      using \<open>\<forall>\<^sub>F w in at z. f w = 0\<close> unfolding laurent_expansion_def by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
    ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
    case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
    define df where "df=zor_poly (\<lambda>x. f (z + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
    define g where "g=(\<lambda>u. u-z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   970
    have "fps_expansion df 0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
        =  fps_expansion (df o g) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
    proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   973
      have "\<exists>\<^sub>F w in at 0. f (z + w) \<noteq> 0" using True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   974
        by (smt (verit, best) add.commute eventually_at_to_0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
            eventually_mono not_frequently)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   976
      from zorder_exist[OF iso ness this,folded df_def]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
      obtain r where "r>0" and df_holo:"df holomorphic_on cball 0 r" and "df 0 \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
          "\<forall>w\<in>cball 0 r - {0}.
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   979
             f (z + w) = df w * w powi (zorder (\<lambda>w. f (z + w)) 0) \<and>
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   980
             df w \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
        by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
      then have df_nz:"\<forall>w\<in>ball 0 r. df w\<noteq>0" by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
      have "(deriv ^^ n) df 0 =  (deriv ^^ n) (df \<circ> g) z" for n
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
        unfolding comp_def g_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
      proof (subst higher_deriv_compose_linear'[where u=1 and c="-z",simplified])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
        show "df holomorphic_on ball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
          using df_holo by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
        show "open (ball z r)" "open (ball 0 r)" "z \<in> ball z r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
          using \<open>r>0\<close> by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
        show " \<And>w. w \<in> ball z r \<Longrightarrow> w - z \<in> ball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
          by (simp add: dist_norm)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
      qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
      then show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
        unfolding fps_expansion_def by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   997
    also have "... = fps_expansion (zor_poly f z) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   998
    proof (rule fps_expansion_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
      have "\<forall>\<^sub>F w in nhds z. zor_poly f z w
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1000
                = zor_poly (\<lambda>u. f (z + u)) 0 (w - z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1001
        apply (rule zor_poly_shift)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
        using True assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
      then show "\<forall>\<^sub>F w in nhds z. (df \<circ> g) w = zor_poly f z w"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
        unfolding df_def g_def comp_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
        by (auto elim:eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1007
    finally show ?thesis unfolding df_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
      by (auto simp: laurent_expansion_def at_to_0[of z]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1009
          eventually_filtermap add_ac zorder_shift')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1010
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1011
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1012
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1014
lemma has_fps_expansion_to_laurent:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1015
  "f has_fps_expansion F \<longleftrightarrow> f has_laurent_expansion fps_to_fls F \<and> f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
proof safe
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
  assume *: "f has_laurent_expansion fps_to_fls F" "f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
  have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
    using * by (intro eventually_nhds_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
  moreover have "eventually (\<lambda>z. z \<noteq> 0 \<longrightarrow> eval_fls (fps_to_fls F) z = f z) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
    using * by (auto simp: has_laurent_expansion_def eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
  ultimately have "eventually (\<lambda>z. f z = eval_fps F z) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
    by eventually_elim
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
       (auto simp: has_laurent_expansion_def eventually_at_filter eval_fps_at_0 eval_fps_to_fls *(2))
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
  thus "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
    using * by (auto simp: has_fps_expansion_def has_laurent_expansion_def eq_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
  assume "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
  thus "f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
    by (metis eval_fps_at_0 has_fps_expansion_imp_holomorphic)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
qed (auto intro: has_laurent_expansion_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1032
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
lemma eval_fps_fls_base_factor [simp]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
  assumes "z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1035
  shows   "eval_fps (fls_base_factor_to_fps F) z = eval_fls F z * z powi -fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1036
  using assms unfolding eval_fls_def by (simp add: power_int_minus field_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1037
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1038
lemma has_fps_expansion_imp_analytic_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1039
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1040
  shows   "f analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1041
  by (meson analytic_at_two assms has_fps_expansion_imp_holomorphic)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1042
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1043
lemma has_fps_expansion_imp_analytic:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1044
  assumes "(\<lambda>x. f (z + x)) has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
  shows   "f analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1046
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
  have "(\<lambda>x. f (z + x)) analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
    by (rule has_fps_expansion_imp_analytic_0) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
  hence "(\<lambda>x. f (z + x)) \<circ> (\<lambda>x. x - z) analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
    by (intro analytic_on_compose_gen analytic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
    by (simp add: o_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
lemma is_pole_cong_asymp_equiv:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
  assumes "f \<sim>[at z] g" "z = z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
  shows   "is_pole f z = is_pole g z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
  using asymp_equiv_at_infinity_transfer[OF assms(1)]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
        asymp_equiv_at_infinity_transfer[OF asymp_equiv_symI[OF assms(1)]] assms(2)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
  unfolding is_pole_def by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
lemma not_is_pole_const [simp]: "\<not>is_pole (\<lambda>_::'a::perfect_space. c :: complex) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
  using not_tendsto_and_filterlim_at_infinity[of "at z" "\<lambda>_::'a. c" c] by (auto simp: is_pole_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
lemma has_laurent_expansion_imp_is_pole_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
  assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
  shows   "is_pole f z \<longleftrightarrow> fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
  assume pole: "is_pole f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
  have [simp]: "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
  proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
    assume "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
    hence "is_pole f z \<longleftrightarrow> is_pole (\<lambda>_. 0 :: complex) z" using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
      by (intro is_pole_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
         (auto simp: has_laurent_expansion_def at_to_0[of z] eventually_filtermap add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
    with pole show False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
      by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1078
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1079
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1080
  note pole
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1081
  also have "is_pole f z \<longleftrightarrow>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1082
             is_pole (\<lambda>w. fls_nth F (fls_subdegree F) * (w - z) powi fls_subdegree F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
    using has_laurent_expansion_imp_asymp_equiv[OF F] by (intro is_pole_cong_asymp_equiv refl)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
  also have "\<dots> \<longleftrightarrow> is_pole (\<lambda>w. (w - z) powi fls_subdegree F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
  finally have pole': \<dots> .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
  have False if "fls_subdegree F \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
    have "(\<lambda>w. (w - z) powi fls_subdegree F) holomorphic_on UNIV"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
      using that by (intro holomorphic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
    hence "\<not>is_pole (\<lambda>w. (w - z) powi fls_subdegree F) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
      by (meson UNIV_I not_is_pole_holomorphic open_UNIV)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1094
    with pole' show False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1095
      by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
  thus "fls_subdegree F < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1098
    by force
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1099
qed (use has_laurent_expansion_imp_is_pole[OF assms] in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1100
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
lemma analytic_at_imp_has_fps_expansion_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
  assumes "f analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1103
  shows   "f has_fps_expansion fps_expansion f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
  using assms has_fps_expansion_fps_expansion analytic_at by fast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
lemma analytic_at_imp_has_fps_expansion:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
  assumes "f analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
  shows   "(\<lambda>x. f (z + x)) has_fps_expansion fps_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
  have "f \<circ> (\<lambda>x. z + x) analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1111
    by (intro analytic_on_compose_gen[OF _ assms] analytic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1112
  hence "(f \<circ> (\<lambda>x. z + x)) has_fps_expansion fps_expansion (f \<circ> (\<lambda>x. z + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1113
    unfolding o_def by (intro analytic_at_imp_has_fps_expansion_0) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1114
  also have "\<dots> = fps_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1115
    by (simp add: fps_expansion_def higher_deriv_shift_0')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1116
  finally show ?thesis by (simp add: add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1117
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1118
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1119
lemma has_laurent_expansion_zorder_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1120
  assumes "f has_laurent_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1121
  shows   "zorder f 0 = fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1123
  define G where "G = fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
  from assms obtain A where A: "0 \<in> A" "open A" "\<And>x. x \<in> A - {0} \<Longrightarrow> eval_fls F x = f x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
    unfolding has_laurent_expansion_def eventually_at_filter eventually_nhds
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
    by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1128
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
  proof (rule zorder_eqI)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
    show "open (A \<inter> eball 0 (fls_conv_radius F))" "0 \<in> A \<inter>  eball 0 (fls_conv_radius F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
      using assms A by (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1132
    show "eval_fps G holomorphic_on A \<inter> eball 0 (fls_conv_radius F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
      by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef G_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
    show "eval_fps G 0 \<noteq> 0" using \<open>F \<noteq> 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
      by (auto simp: eval_fps_at_0 G_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1137
    fix w :: complex assume "w \<in> A \<inter> eball 0 (fls_conv_radius F)" "w \<noteq> 0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1138
    thus "f w = eval_fps G w * (w - 0) powi (fls_subdegree F)"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
      using A unfolding G_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
      by (subst eval_fps_fls_base_factor)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
         (auto simp: complex_powr_of_int power_int_minus field_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
lemma has_laurent_expansion_zorder:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
  assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
  shows   "zorder f z = fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
  using has_laurent_expansion_zorder_0[OF assms] by (simp add: zorder_shift' add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
lemma has_fps_expansion_zorder_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
  assumes "f has_fps_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
  shows   "zorder f 0 = int (subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
  using assms has_laurent_expansion_zorder_0[of f "fps_to_fls F"]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
  by (auto simp: has_fps_expansion_to_laurent fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
lemma has_fps_expansion_zorder:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
  assumes "(\<lambda>w. f (z + w)) has_fps_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1158
  shows   "zorder f z = int (subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
  using has_fps_expansion_zorder_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
  by (simp add: zorder_shift' add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
lemma has_fps_expansion_fls_base_factor_to_fps:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1164
  defines "n \<equiv> fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1165
  defines "c \<equiv> fps_nth (fls_base_factor_to_fps F) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1166
  shows   "(\<lambda>z. if z = 0 then c else f z * z powi -n) has_fps_expansion fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1168
  have "(\<lambda>z. f z * z powi -n) has_laurent_expansion fls_shift (-(-n)) F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
    by (intro laurent_expansion_intros assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1170
  also have "fls_shift (-(-n)) F = fps_to_fls (fls_base_factor_to_fps F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
    by (simp add: n_def fls_shift_nonneg_subdegree)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1172
  also have "(\<lambda>z. f z * z powi - n) has_laurent_expansion fps_to_fls (fls_base_factor_to_fps F) \<longleftrightarrow>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1173
             (\<lambda>z. if z = 0 then c else f z * z powi -n) has_laurent_expansion fps_to_fls (fls_base_factor_to_fps F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
    by (intro has_laurent_expansion_cong) (auto simp: eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
  also have "\<dots> \<longleftrightarrow> (\<lambda>z. if z = 0 then c else f z * z powi -n) has_fps_expansion fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
    by (subst has_fps_expansion_to_laurent) (auto simp: c_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1179
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1180
lemma zero_has_laurent_expansion_imp_eq_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
  assumes "(\<lambda>_. 0) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1182
  shows   "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
  have "at (0 :: complex) \<noteq> bot"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1186
  moreover have "(\<lambda>z. if z = 0 then fls_nth F (fls_subdegree F) else 0) has_fps_expansion
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
          fls_base_factor_to_fps F" (is "?f has_fps_expansion _")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
    using has_fps_expansion_fls_base_factor_to_fps[OF assms] by (simp cong: if_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1189
  hence "isCont ?f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1190
    using has_fps_expansion_imp_continuous by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1191
  hence "?f \<midarrow>0\<rightarrow> fls_nth F (fls_subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
    by (auto simp: isCont_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1193
  moreover have "?f \<midarrow>0\<rightarrow> 0 \<longleftrightarrow> (\<lambda>_::complex. 0 :: complex) \<midarrow>0\<rightarrow> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1194
    by (intro filterlim_cong) (auto simp: eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1195
  hence "?f \<midarrow>0\<rightarrow> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1196
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
  ultimately have "fls_nth F (fls_subdegree F) = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1198
    by (rule tendsto_unique)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1199
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1200
    by (meson nth_fls_subdegree_nonzero)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1203
lemma has_laurent_expansion_unique:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
  assumes "f has_laurent_expansion F" "f has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
  shows   "F = G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1206
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1207
  from assms have "(\<lambda>x. f x - f x) has_laurent_expansion F - G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1208
    by (intro laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1209
  hence "(\<lambda>_. 0) has_laurent_expansion F - G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
  hence "F - G = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
    by (rule zero_has_laurent_expansion_imp_eq_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
lemma laurent_expansion_eqI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
  assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
  shows   "laurent_expansion f z = F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
  using assms has_laurent_expansion_isolated has_laurent_expansion_not_essential
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
        has_laurent_expansion_unique not_essential_has_laurent_expansion by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
lemma laurent_expansion_0_eqI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1224
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1225
  shows   "laurent_expansion f 0 = F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1226
  using assms laurent_expansion_eqI[of f 0] by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1227
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1228
lemma has_laurent_expansion_nonzero_imp_eventually_nonzero:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1229
  assumes "f has_laurent_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
  shows   "eventually (\<lambda>x. f x \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
proof (rule ccontr)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1232
  assume "\<not>eventually (\<lambda>x. f x \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
  with assms have "eventually (\<lambda>x. f x = 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
    by (intro not_essential_frequently_0_imp_eventually_0 has_laurent_expansion_isolated
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
              has_laurent_expansion_not_essential)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
       (auto simp: frequently_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
  hence "(f has_laurent_expansion 0) \<longleftrightarrow> ((\<lambda>_. 0) has_laurent_expansion 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1238
    by (intro has_laurent_expansion_cong) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
  hence "f has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1240
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
  with assms(1) have "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
    using has_laurent_expansion_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1243
  with \<open>F \<noteq> 0\<close> show False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1244
    by contradiction
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1245
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
lemma has_laurent_expansion_eventually_nonzero_iff':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
  shows   "eventually (\<lambda>x. f x \<noteq> 0) (at 0) \<longleftrightarrow> F \<noteq> 0 "
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
  assume "\<forall>\<^sub>F x in at 0. f x \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
  moreover have "\<not> (\<forall>\<^sub>F x in at 0. f x \<noteq> 0)" if "F=0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
    have "\<forall>\<^sub>F x in at 0. f x = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
      using assms that unfolding has_laurent_expansion_def by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
    then show ?thesis unfolding not_eventually
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
      by (auto elim:eventually_frequentlyE)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
  ultimately show "F \<noteq> 0" by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
qed (simp add:has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
lemma has_laurent_expansion_eventually_nonzero_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
  assumes "(\<lambda>w. f (z+w)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1264
  shows   "eventually (\<lambda>x. f x \<noteq> 0) (at z)  \<longleftrightarrow> F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
  apply (subst eventually_at_to_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
  apply (rule has_laurent_expansion_eventually_nonzero_iff')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
  using assms by (simp add:add.commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
lemma has_laurent_expansion_inverse [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
  shows   "(\<lambda>x. inverse (f x)) has_laurent_expansion inverse F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
proof (cases "F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
  thus ?thesis using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
    by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
  define G where "G = laurent_expansion (\<lambda>x. inverse (f x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
  from False have ev: "eventually (\<lambda>z. f z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
    by (intro has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
  have *: "(\<lambda>x. inverse (f x)) has_laurent_expansion G" unfolding G_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
    by (intro not_essential_has_laurent_expansion_0 isolated_singularity_at_inverse not_essential_inverse
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
              has_laurent_expansion_isolated_0[OF assms] has_laurent_expansion_not_essential_0[OF assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
  have "(\<lambda>x. f x * inverse (f x)) has_laurent_expansion F * G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1286
    by (intro laurent_expansion_intros assms *)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
  also have "?this \<longleftrightarrow> (\<lambda>x. 1) has_laurent_expansion F * G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
    by (intro has_laurent_expansion_cong refl eventually_mono[OF ev]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
  finally have "(\<lambda>_. 1) has_laurent_expansion F * G" .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
  moreover have "(\<lambda>_. 1) has_laurent_expansion 1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
  ultimately have "F * G = 1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
    using has_laurent_expansion_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1294
  hence "G = inverse F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
    using inverse_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
  with * show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
lemma has_laurent_expansion_power_int [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
  "f has_laurent_expansion F \<Longrightarrow> (\<lambda>x. f x powi n) has_laurent_expansion (F powi n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
  by (auto simp: power_int_def intro!: laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
lemma has_fps_expansion_0_analytic_continuation:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
  assumes "f has_fps_expansion 0" "f holomorphic_on A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
  assumes "open A" "connected A" "0 \<in> A" "x \<in> A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
  shows   "f x = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
  have "eventually (\<lambda>z. z \<in> A \<and> f z = 0) (nhds 0)" using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
    by (intro eventually_conj eventually_nhds_in_open) (auto simp: has_fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
  then obtain B where B: "open B" "0 \<in> B" "\<forall>z\<in>B. z \<in> A \<and> f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
    unfolding eventually_nhds by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
  proof (rule analytic_continuation_open[where f = f and g = "\<lambda>_. 0"])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
    show "B \<noteq> {}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
      using \<open>open B\<close> B by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
    show "connected A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
      using assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
  qed (use assms B in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
lemma has_laurent_expansion_0_analytic_continuation:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1324
  assumes "f has_laurent_expansion 0" "f holomorphic_on A - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
  assumes "open A" "connected A" "0 \<in> A" "x \<in> A - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
  shows   "f x = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1327
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
  have "eventually (\<lambda>z. z \<in> A - {0} \<and> f z = 0) (at 0)" using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1329
    by (intro eventually_conj eventually_at_in_open) (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1330
  then obtain B where B: "open B" "0 \<in> B" "\<forall>z\<in>B - {0}. z \<in> A - {0} \<and> f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
    unfolding eventually_at_filter eventually_nhds by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
  proof (rule analytic_continuation_open[where f = f and g = "\<lambda>_. 0"])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
    show "B - {0} \<noteq> {}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
      using \<open>open B\<close> \<open>0 \<in> B\<close> by (metis insert_Diff not_open_singleton)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1336
    show "connected (A - {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
      using assms by (intro connected_open_delete) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1338
  qed (use assms B in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
lemma has_fps_expansion_cong:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1342
  assumes "eventually (\<lambda>x. f x = g x) (nhds 0)" "F = G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
  shows   "f has_fps_expansion F \<longleftrightarrow> g has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1344
  using assms(2) by (auto simp: has_fps_expansion_def elim!: eventually_elim2[OF assms(1)])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
lemma zor_poly_has_fps_expansion:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
  assumes "f has_laurent_expansion F" "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
  shows   "zor_poly f 0 has_fps_expansion fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
  note [simp] = \<open>F \<noteq> 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1351
  have "eventually (\<lambda>z. f z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
    by (rule has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
  hence freq: "frequently (\<lambda>z. f z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
    by (rule eventually_frequently[rotated]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
  have *: "isolated_singularity_at f 0" "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
    using has_laurent_expansion_isolated_0[OF assms(1)] has_laurent_expansion_not_essential_0[OF assms(1)]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
  define G where "G = fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
  define n where "n = zorder f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1362
  have n_altdef: "n = fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1363
    using has_laurent_expansion_zorder_0 [OF assms(1)] by (simp add: n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1364
  obtain r where r: "zor_poly f 0 0 \<noteq> 0" "zor_poly f 0 holomorphic_on cball 0 r" "r > 0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1365
                    "\<forall>w\<in>cball 0 r - {0}. f w = zor_poly f 0 w * w powi n \<and>
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
                                         zor_poly f 0 w \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
    using zorder_exist[OF * freq] unfolding n_def by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1368
  obtain r' where r': "r' > 0" "\<forall>x\<in>ball 0 r'-{0}. eval_fls F x = f x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1369
    using assms(1) unfolding has_laurent_expansion_def eventually_at_filter eventually_nhds_metric ball_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1370
    by (auto simp: dist_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1371
  have holo: "zor_poly f 0 holomorphic_on ball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1372
    by (rule holomorphic_on_subset[OF r(2)]) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
  have "(\<lambda>z. if z = 0 then fps_nth G 0 else f z * z powi -n) has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
    unfolding G_def n_altdef by (intro has_fps_expansion_fls_base_factor_to_fps assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
  also have "?this \<longleftrightarrow> zor_poly f 0 has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1377
  proof (intro has_fps_expansion_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1378
    have "eventually (\<lambda>z. z \<in> ball 0 (min r r')) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1379
      using \<open>r > 0\<close> \<open>r' > 0\<close> by (intro eventually_nhds_in_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1380
    thus "\<forall>\<^sub>F x in nhds 0. (if x = 0 then G $ 0 else f x * x powi - n) = zor_poly f 0 x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1381
    proof eventually_elim
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1382
      case (elim w)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1383
      have w: "w \<in> ball 0 r" "w \<in> ball 0 r'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1384
        using elim by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1385
      show ?case
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1386
      proof (cases "w = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1387
        case False
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1388
        hence "f w = zor_poly f 0 w * w powi n"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1389
          using r w by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1390
        thus ?thesis using False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1391
          by (simp add: powr_minus complex_powr_of_int power_int_minus)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1392
      next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1393
        case [simp]: True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
        obtain R where R: "R > 0" "R \<le> r" "R \<le> r'" "R \<le> fls_conv_radius F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
          using \<open>r > 0\<close> \<open>r' > 0\<close> assms(1) unfolding has_laurent_expansion_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
          by (smt (verit, ccfv_SIG) ereal_dense2 ereal_less(2) less_ereal.simps(1) order.strict_implies_order order_trans)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
        have "eval_fps G 0 = zor_poly f 0 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1398
        proof (rule analytic_continuation_open[where f = "eval_fps G" and g = "zor_poly f 0"])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1399
          show "connected (ball 0 R :: complex set)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1400
            by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
          have "of_real R / 2 \<in> ball 0 R - {0 :: complex}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
            using R by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
          thus "ball 0 R - {0 :: complex} \<noteq> {}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1404
            by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1405
          show "eval_fps G holomorphic_on ball 0 R"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1406
            using R less_le_trans[OF _ R(4)] unfolding G_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1407
            by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1408
          show "zor_poly f 0 holomorphic_on ball 0 R"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1409
            by (rule holomorphic_on_subset[OF holo]) (use R in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1410
          show "eval_fps G z = zor_poly f 0 z" if "z \<in> ball 0 R - {0}" for z
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
            using that r r' R n_altdef unfolding G_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
            by (subst eval_fps_fls_base_factor)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1413
               (auto simp: complex_powr_of_int field_simps power_int_minus n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1414
        qed (use R in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
        hence "zor_poly f 0 0 = fps_nth G 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
          by (simp add: eval_fps_at_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
        thus ?thesis by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
      qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
  qed (use r' in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
  finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
    by (simp add: G_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
lemma zorder_geI_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
  assumes "f analytic_on {0}" "f holomorphic_on A" "open A" "connected A" "0 \<in> A" "z \<in> A" "f z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1427
  assumes "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
  shows   "zorder f 0 \<ge> n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
  define F where "F = fps_expansion f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
  from assms have "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
    unfolding F_def using analytic_at_imp_has_fps_expansion_0 by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
  hence laurent: "f has_laurent_expansion fps_to_fls F" and [simp]: "f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
    by (simp_all add: has_fps_expansion_to_laurent)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1435
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
  have [simp]: "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1437
  proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
    assume [simp]: "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
    hence "f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
    proof (cases "z = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
      case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
      have "f has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1443
        using laurent by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1444
      thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
      proof (rule has_laurent_expansion_0_analytic_continuation)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
        show "f holomorphic_on A - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1447
          using assms(2) by (rule holomorphic_on_subset) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1448
      qed (use assms False in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1449
    qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1450
    with \<open>f z \<noteq> 0\<close> show False by contradiction
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1451
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1452
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1453
  have "zorder f 0 = int (subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1454
    using has_laurent_expansion_zorder_0[OF laurent] by (simp add: fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1455
  also have "subdegree F \<ge> n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1456
    using assms by (intro subdegree_geI \<open>F \<noteq> 0\<close>) (auto simp: F_def fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1457
  hence "int (subdegree F) \<ge> int n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1462
lemma zorder_geI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1463
  assumes "f analytic_on {x}" "f holomorphic_on A" "open A" "connected A" "x \<in> A" "z \<in> A" "f z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1464
  assumes "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f x = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1465
  shows   "zorder f x \<ge> n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1466
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1467
  have "zorder f x = zorder (f \<circ> (\<lambda>u. u + x)) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1468
    by (subst zorder_shift) (auto simp: o_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1469
  also have "\<dots> \<ge> n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1470
  proof (rule zorder_geI_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1471
    show "(f \<circ> (\<lambda>u. u + x)) analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1472
      by (intro analytic_on_compose_gen[OF _ assms(1)] analytic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1473
    show "f \<circ> (\<lambda>u. u + x) holomorphic_on ((+) (-x)) ` A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1474
      by (intro holomorphic_on_compose_gen[OF _ assms(2)] holomorphic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1475
    show "connected ((+) (- x) ` A)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1476
      by (intro connected_continuous_image continuous_intros assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1477
    show "open ((+) (- x) ` A)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1478
      by (intro open_translation assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1479
    show "z - x \<in> (+) (- x) ` A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1480
      using \<open>z \<in> A\<close> by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1481
    show "0 \<in> (+) (- x) ` A"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1482
      using \<open>x \<in> A\<close> by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1483
    show "(f \<circ> (\<lambda>u. u + x)) (z - x) \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1484
      using \<open>f z \<noteq> 0\<close> by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1485
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1486
    fix k :: nat assume "k < n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1487
    hence "(deriv ^^ k) f x = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1488
      using assms by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1489
    also have "(deriv ^^ k) f x = (deriv ^^ k) (f \<circ> (+) x) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1490
      by (subst higher_deriv_shift_0) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1491
    finally show "(deriv ^^ k) (f \<circ> (\<lambda>u. u + x)) 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1492
      by (subst add.commute) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1493
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1494
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1495
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1496
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1497
lemma has_laurent_expansion_divide [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1498
  assumes "f has_laurent_expansion F" and "g has_laurent_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1499
  shows   "(\<lambda>x. f x / g x) has_laurent_expansion (F / G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1500
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1501
  have "(\<lambda>x. f x * inverse (g x)) has_laurent_expansion (F * inverse G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1502
    by (intro laurent_expansion_intros assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1504
    by (simp add: field_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1506
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1507
lemma has_laurent_expansion_residue_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1508
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
  shows   "residue f 0 = fls_residue F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
proof (cases "fls_subdegree F \<ge> 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1511
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1512
  have "residue f 0 = residue (eval_fls F) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1513
    using assms by (intro residue_cong) (auto simp: has_laurent_expansion_def eq_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1514
  also have "\<dots> = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1515
    by (rule residue_holo[OF _ _ holomorphic_on_eval_fls[OF order.refl]])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1516
       (use True assms in \<open>auto simp: has_laurent_expansion_def zero_ereal_def\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1517
  also have "\<dots> = fls_residue F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1518
    using True by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1519
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1520
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1521
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1522
  hence "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
  have *: "zor_poly f 0 has_fps_expansion fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
    by (intro zor_poly_has_fps_expansion False assms \<open>F \<noteq> 0\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1527
  have "residue f 0 = (deriv ^^ (nat (-zorder f 0) - 1)) (zor_poly f 0) 0 / fact (nat (- zorder f 0) - 1)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1528
    by (intro residue_pole_order has_laurent_expansion_isolated_0[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1529
              has_laurent_expansion_imp_is_pole_0[OF assms]) (use False in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1530
  also have "\<dots> = fls_residue F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1531
    using has_laurent_expansion_zorder_0[OF assms \<open>F \<noteq> 0\<close>] False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1532
    by (subst fps_nth_fps_expansion [OF *, symmetric]) (auto simp: of_nat_diff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1533
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1536
lemma has_laurent_expansion_residue:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1537
  assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1538
  shows   "residue f z = fls_residue F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
  using has_laurent_expansion_residue_0[OF assms] by (simp add: residue_shift_0')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1541
lemma eval_fls_has_laurent_expansion [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1542
  assumes "fls_conv_radius F > 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1543
  shows   "eval_fls F has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
  using assms by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1546
lemma fps_expansion_unique_complex:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1547
  fixes F G :: "complex fps"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1548
  assumes "f has_fps_expansion F" "f has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
  shows   "F = G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
  using assms unfolding fps_eq_iff by (auto simp: fps_eq_iff fps_nth_fps_expansion)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1551
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1552
lemma fps_expansion_eqI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1553
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
  shows   "fps_expansion f 0 = F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
  using assms unfolding fps_eq_iff
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1556
  by (auto simp: fps_eq_iff fps_nth_fps_expansion fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1557
82541
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1558
lemma holomorphic_on_imp_fps_conv_radius_ge:
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1559
  assumes "f has_fps_expansion F" "f holomorphic_on eball 0 r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1560
  shows   "fps_conv_radius F \<ge> r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1561
proof -
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1562
  define n where "n = subdegree F"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1563
  have "fps_conv_radius (fps_expansion f 0) \<ge> r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1564
    by (intro conv_radius_fps_expansion assms)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1565
  also have "fps_expansion f 0 = F"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1566
    using assms by (intro fps_expansion_eqI)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1567
  finally show ?thesis
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1568
    by simp
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1569
qed
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1570
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
lemma has_fps_expansion_imp_eval_fps_eq:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1572
  assumes "f has_fps_expansion F" "norm z < r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1573
  assumes "f holomorphic_on ball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1574
  shows   "eval_fps F z = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1575
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
  have [simp]: "fps_expansion f 0 = F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
    by (rule fps_expansion_eqI) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
  have *: "f holomorphic_on eball 0 (ereal r)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
    using assms by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1580
  from conv_radius_fps_expansion[OF *] have "fps_conv_radius F \<ge> ereal r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1581
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
  have "eval_fps (fps_expansion f 0) z = f (0 + z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
    by (rule eval_fps_expansion'[OF *]) (use assms in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1584
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
82541
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1588
lemma has_fps_expansion_imp_sums_complex:
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1589
  fixes F :: "complex fps"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1590
  assumes "f has_fps_expansion F" "f holomorphic_on eball 0 r" "ereal (norm z) < r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1591
  shows   "(\<lambda>n. fps_nth F n * z ^ n) sums f z"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1592
proof -
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1593
  have r: "fps_conv_radius F \<ge> r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1594
    using assms(1,2) by (rule holomorphic_on_imp_fps_conv_radius_ge)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1595
  from assms obtain R where R: "norm z < R" "ereal R < r"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1596
    using ereal_dense2 less_ereal.simps(1) by blast
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1597
  have z: "norm z < fps_conv_radius F"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1598
    using r R assms(3) by order
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1599
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1600
  have "summable (\<lambda>n. fps_nth F n * z ^ n)"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1601
    by (rule summable_fps) (use z in auto)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1602
  moreover have "eval_fps F z = f z"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1603
  proof (rule has_fps_expansion_imp_eval_fps_eq[where r = R])
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1604
    have *: "ereal (norm z) < r" if "norm z < R" for z :: complex
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1605
      using that R ereal_le_less less_imp_le by blast
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1606
    show "f holomorphic_on ball 0 R"
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1607
      using assms(2) by (rule holomorphic_on_subset) (use * in auto)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1608
  qed (use R assms(1) in auto)
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1609
  ultimately show ?thesis
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1610
    unfolding eval_fps_def sums_iff by simp
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1611
qed
5160b68e78a9 some material on power series and infinite products
Manuel Eberl <eberlm@in.tum.de>
parents: 82530
diff changeset
  1612
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
lemma fls_conv_radius_ge:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
  assumes "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
  assumes "f holomorphic_on eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
  shows   "fls_conv_radius F \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
  define n where "n = fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
  define G where "G = fls_base_factor_to_fps F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
  define g where "g = (\<lambda>z. if z = 0 then fps_nth G 0 else f z * z powi -n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
  have G: "g has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
    unfolding G_def g_def n_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
    by (intro has_fps_expansion_fls_base_factor_to_fps assms)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
  have "(\<lambda>z. f z * z powi -n) holomorphic_on eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
    by (intro holomorphic_intros assms) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
  also have "?this \<longleftrightarrow> g holomorphic_on eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
    by (intro holomorphic_cong) (auto simp: g_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
  finally have "g analytic_on eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
    by (subst analytic_on_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
  moreover have "g analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
    using G has_fps_expansion_imp_analytic_0 by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
  ultimately have "g analytic_on (eball 0 r - {0} \<union> {0})"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
    by (subst analytic_on_Un) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
  hence "g analytic_on eball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1635
    by (rule analytic_on_subset) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
  hence "g holomorphic_on eball 0 r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
    by (subst (asm) analytic_on_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
  hence "fps_conv_radius (fps_expansion g 0) \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
    by (intro conv_radius_fps_expansion)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
  also have "fps_expansion g 0 = G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
    using G by (intro fps_expansion_eqI)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
  finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
    by (simp add: fls_conv_radius_altdef G_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
lemma eval_fls_eqI:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
  assumes "f has_laurent_expansion F" "f holomorphic_on eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
  assumes "z \<in> eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
  shows   "eval_fls F z = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
  have conv: "fls_conv_radius F \<ge> r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
    by (intro fls_conv_radius_ge[OF assms(1,2)])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
  have "(\<lambda>z. eval_fls F z - f z) has_laurent_expansion F - F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1654
    using assms by (intro laurent_expansion_intros assms) (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
  hence "(\<lambda>z. eval_fls F z - f z) has_laurent_expansion 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
  hence "eval_fls F z - f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1658
  proof (rule has_laurent_expansion_0_analytic_continuation)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
    have "ereal 0 \<le> ereal (norm z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
      by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
    also have "norm z < r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1662
      using assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
    finally have "r > 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
      by (simp add: zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
    thus "open (eball 0 r :: complex set)" "connected (eball 0 r :: complex set)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
         "0 \<in> eball 0 r" "z \<in> eball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
      using assms by (auto simp: zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
  qed (auto intro!: holomorphic_intros assms less_le_trans[OF _ conv] split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
  thus ?thesis by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1671
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
lemma fls_nth_as_contour_integral:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
  assumes F: "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1674
  assumes holo: "f holomorphic_on ball 0 r - {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1675
  assumes R: "0 < R" "R < r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1676
  shows "((\<lambda>z. f z * z powi (-(n+1))) has_contour_integral
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
            complex_of_real (2 * pi) * \<i> * fls_nth F n) (circlepath 0 R)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
  define I where "I = (\<lambda>z. f z * z powi (-(n+1)))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1680
  have "(I has_contour_integral complex_of_real (2 * pi) * \<i> * residue I 0) (circlepath 0 R)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
  proof (rule base_residue)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
    show "open (ball (0::complex) r)" "0 \<in> ball (0::complex) r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1683
      using R F by (auto simp: has_laurent_expansion_def zero_ereal_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
  qed (use R in \<open>auto intro!: holomorphic_intros holomorphic_on_subset[OF holo]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1685
                      simp: I_def split: if_splits\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
  also have "residue I 0 = fls_residue (fls_shift (n + 1) F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
    unfolding I_def by (intro has_laurent_expansion_residue_0 laurent_expansion_intros F)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
  also have "\<dots> = fls_nth F n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1689
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1690
  finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1691
    by (simp add: I_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1692
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1693
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
lemma tendsto_0_subdegree_iff_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
  assumes F:"f has_laurent_expansion F" and "F\<noteq>0"
82529
ff4b062aae57 moved some lemmas to where they fit better
Manuel Eberl <manuel@pruvisto.org>
parents: 82517
diff changeset
  1696
  shows "(f \<midarrow>0\<rightarrow> 0) \<longleftrightarrow> fls_subdegree F > 0"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1698
  have ?thesis if "is_pole f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1700
    have "fls_subdegree F <0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1701
      using is_pole_0_imp_neg_fls_subdegree[OF F that] .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1702
    moreover then have "\<not> f \<midarrow>0\<rightarrow>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
      using \<open>is_pole f 0\<close> F at_neq_bot
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
        has_laurent_expansion_imp_filterlim_infinity_0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
        not_tendsto_and_filterlim_at_infinity that
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
      by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
    ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
  moreover have ?thesis if "\<not>is_pole f 0" "\<exists>x. f \<midarrow>0\<rightarrow>x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1711
    have "fls_subdegree F \<ge>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
      using has_laurent_expansion_imp_is_pole_0[OF F] that(1)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
      by linarith
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
    have "f \<midarrow>0\<rightarrow>0" if "fls_subdegree F > 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
      using fls_eq0_below_subdegree[OF that]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
      by (metis F \<open>0 \<le> fls_subdegree F\<close> has_laurent_expansion_imp_tendsto_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1717
    moreover have "fls_subdegree F > 0" if "f \<midarrow>0\<rightarrow>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1718
    proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1719
      have False if "fls_subdegree F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1720
      proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
        have "f \<midarrow>0\<rightarrow> fls_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
          using has_laurent_expansion_imp_tendsto_0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
              [OF F \<open>fls_subdegree F \<ge>0\<close>] .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
        then have "fls_nth F 0 = 0" using \<open>f \<midarrow>0\<rightarrow>0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
          using LIM_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
        then have "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
          using nth_fls_subdegree_zero_iff \<open>fls_subdegree F = 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
          by metis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
        with \<open>F\<noteq>0\<close> show False by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
      qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
      with \<open>fls_subdegree F \<ge>0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
      show ?thesis by fastforce
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1734
    ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
  moreover have "is_pole f 0 \<or> (\<exists>x. f \<midarrow>0\<rightarrow>x)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
    have "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1739
      using F has_laurent_expansion_not_essential_0 by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1740
    then show ?thesis unfolding not_essential_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
      by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1743
  ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
lemma tendsto_0_subdegree_iff:
82529
ff4b062aae57 moved some lemmas to where they fit better
Manuel Eberl <manuel@pruvisto.org>
parents: 82517
diff changeset
  1747
  assumes F: "(\<lambda>w. f (z+w)) has_laurent_expansion F" and "F \<noteq> 0"
ff4b062aae57 moved some lemmas to where they fit better
Manuel Eberl <manuel@pruvisto.org>
parents: 82517
diff changeset
  1748
  shows "(f \<midarrow>z\<rightarrow> 0) \<longleftrightarrow> fls_subdegree F > 0"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
  apply (subst Lim_at_zero)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
  apply (rule tendsto_0_subdegree_iff_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
  using assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
lemma is_pole_0_deriv_divide_iff:
82529
ff4b062aae57 moved some lemmas to where they fit better
Manuel Eberl <manuel@pruvisto.org>
parents: 82517
diff changeset
  1754
  assumes F: "f has_laurent_expansion F" and "F \<noteq> 0"
ff4b062aae57 moved some lemmas to where they fit better
Manuel Eberl <manuel@pruvisto.org>
parents: 82517
diff changeset
  1755
  shows "is_pole (\<lambda>x. deriv f x / f x) 0 \<longleftrightarrow> is_pole f 0 \<or> (f \<midarrow>0\<rightarrow> 0)"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
  have "(\<lambda>x. deriv f x / f x) has_laurent_expansion fls_deriv F / F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
    using F by (auto intro:laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
  have "is_pole (\<lambda>x. deriv f x / f x) 0 \<longleftrightarrow>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
            fls_subdegree (fls_deriv F / F) < 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
    apply (rule is_pole_fls_subdegree_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
    using F by (auto intro:laurent_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
  also have "... \<longleftrightarrow> is_pole f 0 \<or> (f \<midarrow>0\<rightarrow>0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
  proof (cases "fls_subdegree F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
    case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
    then have "fls_subdegree (fls_deriv F / F) \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
      by (metis diff_zero div_0 \<open>F\<noteq>0\<close> fls_deriv_subdegree0
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
          fls_divide_subdegree)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
    moreover then have "\<not> is_pole f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
      by (metis F True is_pole_0_imp_neg_fls_subdegree less_le)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
    moreover have "\<not> (f \<midarrow>0\<rightarrow>0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
      using tendsto_0_subdegree_iff_0[OF F \<open>F\<noteq>0\<close>] True by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
    ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
    case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1777
    then have "fls_deriv F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1778
      by (metis fls_const_subdegree fls_deriv_eq_0_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1779
    then have "fls_subdegree (fls_deriv F / F) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1780
              fls_subdegree (fls_deriv F) - fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
      by (rule fls_divide_subdegree[OF _ \<open>F\<noteq>0\<close>])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
    moreover have "fls_subdegree (fls_deriv F) = fls_subdegree F - 1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
      using fls_subdegree_deriv[OF False] .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
    ultimately have "fls_subdegree (fls_deriv F / F) < 0" by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
    moreover have "f \<midarrow>0\<rightarrow> 0 = (0 < fls_subdegree F)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
      using tendsto_0_subdegree_iff_0[OF F \<open>F \<noteq> 0\<close>] .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1787
    moreover have "is_pole f 0 = (fls_subdegree F < 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1788
      using is_pole_fls_subdegree_iff F by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
    ultimately show ?thesis using False by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1793
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1794
lemma is_pole_deriv_divide_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
  assumes F:"(\<lambda>w. f (z+w))  has_laurent_expansion F" and "F\<noteq>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
  shows "is_pole (\<lambda>x. deriv f x / f x) z \<longleftrightarrow> is_pole f z \<or> (f \<midarrow>z\<rightarrow>0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
  define ff df where "ff=(\<lambda>w. f (z+w))" and "df=(\<lambda>w. deriv f (z + w))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
  have "is_pole (\<lambda>x. deriv f x / f x) z
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
          \<longleftrightarrow> is_pole (\<lambda>x. deriv ff x / ff x) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1801
    unfolding ff_def df_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
    by (simp add:deriv_shift_0' is_pole_shift_0' comp_def algebra_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
  moreover have "is_pole f z \<longleftrightarrow> is_pole ff 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
    unfolding ff_def by (auto simp:is_pole_shift_0')
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
  moreover have "(f \<midarrow>z\<rightarrow>0) \<longleftrightarrow> (ff \<midarrow>0\<rightarrow>0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
    unfolding ff_def by (simp add: LIM_offset_zero_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
  moreover have "is_pole (\<lambda>x. deriv ff x / ff x) 0 = (is_pole ff 0 \<or> ff \<midarrow>0\<rightarrow> 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
    apply (rule is_pole_0_deriv_divide_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
    using F ff_def \<open>F\<noteq>0\<close> by blast+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
  ultimately show ?thesis by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
lemma subdegree_imp_eventually_deriv_nzero_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
  assumes F:"f has_laurent_expansion F" and "fls_subdegree F\<noteq>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
  shows "eventually (\<lambda>z. deriv f z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
  have "deriv f has_laurent_expansion fls_deriv F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
    using has_laurent_expansion_deriv[OF F] .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
  moreover have "fls_deriv F\<noteq>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
    using \<open>fls_subdegree F\<noteq>0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
    by (metis fls_const_subdegree fls_deriv_eq_0_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
  ultimately show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
    using has_laurent_expansion_eventually_nonzero_iff' by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
lemma subdegree_imp_eventually_deriv_nzero:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
  assumes F:"(\<lambda>w. f (z+w)) has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
      and "fls_subdegree F\<noteq>0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
  shows "eventually (\<lambda>w. deriv f w \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1831
  have "\<forall>\<^sub>F x in at 0. deriv (\<lambda>w. f (z + w)) x \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
    using subdegree_imp_eventually_deriv_nzero_0 assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
  then show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
    apply (subst eventually_at_to_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1835
    by (simp add:deriv_shift_0' comp_def algebra_simps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
lemma has_fps_expansion_imp_asymp_equiv_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
  fixes f :: "complex \<Rightarrow> complex"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
  assumes F: "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
  defines "n \<equiv> subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
  shows   "f \<sim>[nhds 0] (\<lambda>z. fps_nth F n * z ^ n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
  have F': "f has_laurent_expansion fps_to_fls F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
    using F has_laurent_expansion_fps by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
  have "f \<sim>[at 0] (\<lambda>z. fps_nth F n * z ^ n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
    using has_laurent_expansion_imp_asymp_equiv_0[OF F']
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
    by (simp add: fls_subdegree_fls_to_fps n_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1850
  moreover have "f 0 = fps_nth F n * 0 ^ n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
    using F by (auto simp: n_def has_fps_expansion_to_laurent power_0_left)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1852
  ultimately show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
    by (auto simp: asymp_equiv_nhds_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1856
lemma has_fps_expansion_imp_tendsto_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1857
  fixes f :: "complex \<Rightarrow> complex"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
  shows   "(f \<longlongrightarrow> fps_nth F 0) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
proof (rule asymp_equiv_tendsto_transfer)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
  show "(\<lambda>z. fps_nth F (subdegree F) * z ^ subdegree F) \<sim>[nhds 0] f"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
    by (rule asymp_equiv_symI, rule has_fps_expansion_imp_asymp_equiv_0) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
  have "((\<lambda>z. F $ subdegree F * z ^ subdegree F) \<longlongrightarrow> F $ 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
    by (rule tendsto_eq_intros refl | simp)+ (auto simp: power_0_left)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
  thus "((\<lambda>z. F $ subdegree F * z ^ subdegree F) \<longlongrightarrow> F $ 0) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
    by (auto simp: tendsto_nhds_iff power_0_left)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
lemma has_fps_expansion_imp_0_eq_fps_nth_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1871
  shows   "f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
  have "eventually (\<lambda>x. f x = eval_fps F x) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
    using assms by (auto simp: has_fps_expansion_def eq_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
  then obtain A where "open A" "0 \<in> A" "\<forall>x\<in>A. f x = eval_fps F x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
    unfolding eventually_nhds by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
  hence "f 0 = eval_fps F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
    by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
    by (simp add: eval_fps_at_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1883
lemma fls_nth_compose_aux:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
  assumes "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1885
  assumes G: "g has_fps_expansion G" "fps_nth G 0 = 0" "fps_deriv G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
  assumes "(f \<circ> g) has_laurent_expansion H"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
  shows   "fls_nth H (int n) = fps_nth (fps_compose F G) n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
  using assms(1,5)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
proof (induction n arbitrary: f F H rule: less_induct)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
  case (less n f F H)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
  have [simp]: "g 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
    using has_fps_expansion_imp_0_eq_fps_nth_0[OF G(1)] G(2) by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
  have ana_f: "f analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
    using less.prems by (meson has_fps_expansion_imp_analytic_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
  have ana_g: "g analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
    using G by (meson has_fps_expansion_imp_analytic_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
  have "(f \<circ> g) has_laurent_expansion fps_to_fls (fps_expansion (f \<circ> g) 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
    by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros has_laurent_expansion_fps
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
             analytic_on_compose_gen ana_f ana_g)+ auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1900
  with less.prems have "H = fps_to_fls (fps_expansion (f \<circ> g) 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
    using has_laurent_expansion_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
  also have "fls_subdegree \<dots> \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
    by (simp add: fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
  finally have subdeg: "fls_subdegree H \<ge> 0" .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
  show ?case
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
  proof (cases "n = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
    case [simp]: True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
    have lim_g: "g \<midarrow>0\<rightarrow> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
      using has_laurent_expansion_imp_tendsto_0[of g "fps_to_fls G"] G
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
      by (auto simp: fls_subdegree_fls_to_fps_gt0 has_fps_expansion_to_laurent)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
    have lim_f: "(f \<longlongrightarrow> fps_nth F 0) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
      by (intro has_fps_expansion_imp_tendsto_0 less.prems)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
    have "(\<lambda>x. f (g x)) \<midarrow>0\<rightarrow> fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
      by (rule filterlim_compose[OF lim_f lim_g])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
    moreover have "(f \<circ> g) \<midarrow>0\<rightarrow> fls_nth H 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
      by (intro has_laurent_expansion_imp_tendsto_0 less.prems subdeg)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
    ultimately have "fps_nth F 0 = fls_nth H 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
      using tendsto_unique by (force simp: o_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
      by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
    case n: False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
    define GH where "GH = (fls_deriv H / fls_deriv (fps_to_fls G))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
    define GH' where "GH' = fls_regpart GH"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
    have "(\<lambda>x. deriv (f \<circ> g) x / deriv g x) has_laurent_expansion
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
          fls_deriv H / fls_deriv (fps_to_fls G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1929
      by (intro laurent_expansion_intros less.prems has_laurent_expansion_fps[of _ G] G)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1930
    also have "?this \<longleftrightarrow> (deriv f \<circ> g) has_laurent_expansion fls_deriv H / fls_deriv (fps_to_fls G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
    proof (rule has_laurent_expansion_cong)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
      from ana_f obtain r1 where r1: "r1 > 0" "f holomorphic_on ball 0 r1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
        unfolding analytic_on_def by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
      from ana_g obtain r2 where r2: "r2 > 0" "g holomorphic_on ball 0 r2"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1935
        unfolding analytic_on_def by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1936
      have lim_g: "g \<midarrow>0\<rightarrow> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
        using has_laurent_expansion_imp_tendsto_0[of g "fps_to_fls G"] G
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1938
        by (auto simp: fls_subdegree_fls_to_fps_gt0 has_fps_expansion_to_laurent)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
      moreover have "open (ball 0 r1)" "0 \<in> ball 0 r1"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1940
        using r1 by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
      ultimately have "eventually (\<lambda>x. g x \<in> ball 0 r1) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
        unfolding tendsto_def by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1943
      moreover have "eventually (\<lambda>x. deriv g x \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1944
        using G fps_to_fls_eq_0_iff has_fps_expansion_deriv has_fps_expansion_to_laurent
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
              has_laurent_expansion_nonzero_imp_eventually_nonzero by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1946
      moreover have "eventually (\<lambda>x. x \<in> ball 0 (min r1 r2) - {0}) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
        by (intro eventually_at_in_open) (use r1 r2 in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
      ultimately show "eventually (\<lambda>x. deriv (f \<circ> g) x / deriv g x = (deriv f \<circ> g) x) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1949
      proof eventually_elim
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
        case (elim x)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1951
        thus ?case using r1 r2
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1952
          by (subst deriv_chain)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1953
             (auto simp: field_simps holomorphic_on_def at_within_open[of _ "ball _ _"])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1954
      qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1955
    qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
    finally have GH: "(deriv f \<circ> g) has_laurent_expansion GH"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
      unfolding GH_def .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
    have "(deriv f \<circ> g) has_laurent_expansion fps_to_fls (fps_expansion (deriv f \<circ> g) 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
      by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros has_laurent_expansion_fps
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
               analytic_on_compose_gen ana_f ana_g)+ auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
    with GH have "GH = fps_to_fls (fps_expansion (deriv f \<circ> g) 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
      using has_laurent_expansion_unique by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1964
    also have "fls_subdegree \<dots> \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
      by (simp add: fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
    finally have subdeg': "fls_subdegree GH \<ge> 0" .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
    have "deriv f has_fps_expansion fps_deriv F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
      by (intro fps_expansion_intros less.prems)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
    from this and GH have IH: "fls_nth GH (int k) = fps_nth (fps_compose (fps_deriv F) G) k"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1971
      if "k < n" for k
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
      by (intro less.IH that)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1973
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1974
    have "fps_nth (fps_compose (fps_deriv F) G) n = (\<Sum>i=0..n. of_nat (Suc i) * F $ Suc i * G ^ i $ n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
      by (simp add: fps_compose_nth)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
    have "fps_nth (fps_compose F G) n =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
            fps_nth (fps_deriv (fps_compose F G)) (n - 1) / of_nat n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
      using n by (cases n) (auto simp del: of_nat_Suc)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
    also have "fps_deriv (fps_compose F G) = fps_compose (fps_deriv F) G * fps_deriv G "
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
      using G by (subst fps_compose_deriv) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
    also have "fps_nth \<dots> (n - 1) = (\<Sum>i=0..n-1. (fps_deriv F oo G) $ i * fps_deriv G $ (n - 1 - i))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
      unfolding fps_mult_nth ..
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
    also have "\<dots> = (\<Sum>i=0..n-1. fps_nth GH' i * of_nat (n - i) * G $ (n - i))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
      using n by (intro sum.cong) (auto simp: IH Suc_diff_Suc GH'_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
    also have "\<dots> = (\<Sum>i=0..n. fps_nth GH' i * of_nat (n - i) * G $ (n - i))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
      by (intro sum.mono_neutral_left) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
    also have "\<dots> = fps_nth (GH' * Abs_fps (\<lambda>i. of_nat i * fps_nth G i)) n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1989
      by (simp add: fps_mult_nth mult_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
    also have "Abs_fps (\<lambda>i. of_nat i * fps_nth G i) = fps_X * fps_deriv G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
      by (simp add: fps_mult_fps_X_deriv_shift)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
    also have "fps_nth (GH' * (fps_X * fps_deriv G)) n =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
               fls_nth (fps_to_fls (GH' * (fps_X * fps_deriv G))) (int n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
      by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
    also have "fps_to_fls (GH' * (fps_X * fps_deriv G)) =
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
                 GH * fps_to_fls (fps_deriv G) * fls_X"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
      using subdeg' by (simp add: mult_ac fls_times_fps_to_fls GH'_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
    also have "GH * fps_to_fls (fps_deriv G) = fls_deriv H"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1999
      unfolding GH_def using G  by (simp add: fls_deriv_fps_to_fls)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2000
    also have "fls_deriv H * fls_X = fls_shift (-1) (fls_deriv H)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
      using fls_X_times_conv_shift(2) by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
    finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
      using n by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
lemma has_fps_expansion_compose [fps_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
  fixes f g :: "complex \<Rightarrow> complex"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
  assumes F: "f has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
  assumes G: "g has_fps_expansion G" "fps_nth G 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
  shows   "(f \<circ> g) has_fps_expansion fps_compose F G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
proof (cases "fps_deriv G = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
  case False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
  have [simp]: "g 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
    using has_fps_expansion_imp_0_eq_fps_nth_0[OF G(1)] G(2) False by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
  have ana_f: "f analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
    using F by (meson has_fps_expansion_imp_analytic_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
  have ana_g: "g analytic_on {0}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
    using G by (meson has_fps_expansion_imp_analytic_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
  have fg: "(f \<circ> g) has_fps_expansion fps_expansion (f \<circ> g) 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
    by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
         analytic_on_compose_gen ana_f ana_g)+ auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
  have "fls_nth (fps_to_fls (fps_expansion (f \<circ> g) 0)) (int n) = fps_nth (fps_compose F G) n" for n
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2025
    by (rule fls_nth_compose_aux has_laurent_expansion_fps F G False fg)+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
  hence "fps_expansion (f \<circ> g) 0 = fps_compose F G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
    by (simp add: fps_eq_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
  thus ?thesis using fg
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2031
  case True
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
  have [simp]: "f 0 = fps_nth F 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
    using F by (auto dest: has_fps_expansion_imp_0_eq_fps_nth_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2034
  from True have "fps_nth G n = 0" for n
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
    using G(2) by (cases n) (auto simp del: of_nat_Suc)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
  hence [simp]: "G = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
    by (auto simp: fps_eq_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2038
  have "(\<lambda>_. f 0) has_fps_expansion fps_const (f 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
    by (intro fps_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2040
  also have "eventually (\<lambda>x. g x = 0) (nhds 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
    using G by (auto simp: has_fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
  hence "(\<lambda>_. f 0) has_fps_expansion fps_const (f 0) \<longleftrightarrow> (f \<circ> g) has_fps_expansion fps_const (f 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
    by (intro has_fps_expansion_cong) (auto elim!: eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
  thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
lemma has_fps_expansion_fps_to_fls:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
  assumes "f has_laurent_expansion fps_to_fls F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2051
  shows   "(\<lambda>z. if z = 0 then fps_nth F 0 else f z) has_fps_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
  (is "?f' has_fps_expansion _")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
  have "f has_laurent_expansion fps_to_fls F \<longleftrightarrow> ?f' has_laurent_expansion fps_to_fls F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
    by (intro has_laurent_expansion_cong) (auto simp: eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
  with assms show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
    by (auto simp: has_fps_expansion_to_laurent)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2059
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
lemma has_laurent_expansion_compose [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
  fixes f g :: "complex \<Rightarrow> complex"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
  assumes F: "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
  assumes G: "g has_laurent_expansion fps_to_fls G" "fps_nth G 0 = 0" "G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2064
  shows   "(f \<circ> g) has_laurent_expansion fls_compose_fps F G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
  from assms have lim_g: "g \<midarrow>0\<rightarrow> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2067
    by (subst tendsto_0_subdegree_iff_0[OF G(1)])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2068
       (auto simp: fls_subdegree_fls_to_fps subdegree_pos_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
  have ev1: "eventually (\<lambda>z. g z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
    using \<open>G \<noteq> 0\<close> G(1) fps_to_fls_eq_0_iff has_laurent_expansion_fps
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
           has_laurent_expansion_nonzero_imp_eventually_nonzero by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2072
  moreover have "eventually (\<lambda>z. z \<noteq> 0) (at (0 :: complex))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2073
    by (auto simp: eventually_at_filter)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
  ultimately have ev: "eventually (\<lambda>z. z \<noteq> 0 \<and> g z \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
    by eventually_elim blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
  from ev1 and lim_g have lim_g': "filterlim g (at 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
    by (auto simp: filterlim_at)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
  define g' where "g' = (\<lambda>z. if z = 0 then fps_nth G 0 else g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
  proof (cases "F = 0")
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
    assume [simp]: "F = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
    have "eventually (\<lambda>z. f z = 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
      using F by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
    hence "eventually (\<lambda>z. f (g z) = 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
      using lim_g' by (rule eventually_compose_filterlim)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2087
    thus ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2088
      by (auto simp: has_laurent_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
  next
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
    assume [simp]: "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
    define n where "n = fls_subdegree F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
    define f' where
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
      "f' = (\<lambda>z. if z = 0 then fps_nth (fls_base_factor_to_fps F) 0 else f z * z powi -n)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
    have "((\<lambda>z. (f' \<circ> g') z * g z powi n)) has_laurent_expansion fls_compose_fps F G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
      unfolding f'_def n_def fls_compose_fps_def g'_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
      by (intro fps_expansion_intros laurent_expansion_intros has_fps_expansion_fps_to_fls
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
                has_fps_expansion_fls_base_factor_to_fps assms has_laurent_expansion_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
    also have "?this \<longleftrightarrow> ?thesis"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
      by (intro has_laurent_expansion_cong eventually_mono[OF ev])
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
         (auto simp: f'_def power_int_minus g'_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
    finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
lemma has_laurent_expansion_fls_X_inv [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
  "inverse has_laurent_expansion fls_X_inv"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
  using has_laurent_expansion_inverse[OF has_laurent_expansion_fps_X]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
  by (simp add: fls_inverse_X)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
lemma zorder_times_analytic:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2111
  assumes "f analytic_on {z}" "g analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
  assumes "eventually (\<lambda>z. f z * g z \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
  shows   "zorder (\<lambda>z. f z * g z) z = zorder f z + zorder g z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
  have *: "(\<lambda>w. f (z + w)) has_fps_expansion fps_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
          "(\<lambda>w. g (z + w)) has_fps_expansion fps_expansion g z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
          "(\<lambda>w. f (z + w) * g (z + w)) has_fps_expansion fps_expansion f z * fps_expansion g z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
    by (intro fps_expansion_intros analytic_at_imp_has_fps_expansion assms)+
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
  have [simp]: "fps_expansion f z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
  proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
    assume "fps_expansion f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
    hence "eventually (\<lambda>z. f z * g z = 0) (at z)" using *(1)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
      by (auto simp: has_fps_expansion_0_iff nhds_to_0' eventually_filtermap eventually_at_filter
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
               elim: eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
    with assms(3) have "eventually (\<lambda>z. False) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
      by eventually_elim auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
    thus False by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
  have [simp]: "fps_expansion g z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
  proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
    assume "fps_expansion g z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
    hence "eventually (\<lambda>z. f z * g z = 0) (at z)" using *(2)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
      by (auto simp: has_fps_expansion_0_iff nhds_to_0' eventually_filtermap eventually_at_filter
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
               elim: eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
    with assms(3) have "eventually (\<lambda>z. False) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2136
      by eventually_elim auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2137
    thus False by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2139
  from *[THEN has_fps_expansion_zorder] show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2140
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
lemma zorder_const [simp]: "c \<noteq> 0 \<Longrightarrow> zorder (\<lambda>_. c) z = 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2144
  by (intro zorder_eqI[where S = UNIV]) auto
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
lemma zorder_prod_analytic:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
  assumes "\<And>x. x \<in> A \<Longrightarrow> f x analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
  assumes "eventually (\<lambda>z. (\<Prod>x\<in>A. f x z) \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
  shows   "zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z = (\<Sum>x\<in>A. zorder (f x) z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2150
  using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2151
proof (induction A rule: infinite_finite_induct)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
  case (insert x A)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
  have "zorder (\<lambda>z. f x z * (\<Prod>x\<in>A. f x z)) z = zorder (f x) z + zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
    using insert.prems insert.hyps by (intro zorder_times_analytic analytic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
  also have "zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z = (\<Sum>x\<in>A. zorder (f x) z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
    using insert.prems insert.hyps by (intro insert.IH) (auto elim!: eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
  finally show ?case using insert
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2158
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
qed auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
lemma zorder_eq_0I:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2162
  assumes "g analytic_on {z}" "g z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
  shows   "zorder g z = 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2164
  using analytic_at assms zorder_eqI by fastforce
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
lemma zorder_pos_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
  assumes "f holomorphic_on A" "open A" "z \<in> A" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
  shows   "zorder f z > 0 \<longleftrightarrow> f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
  have "f analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
    using assms analytic_at by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2172
  hence *: "(\<lambda>w. f (z + w)) has_fps_expansion fps_expansion f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2173
    using analytic_at_imp_has_fps_expansion by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2174
  have nz: "fps_expansion f z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
  proof
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
    assume "fps_expansion f z = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
    hence "eventually (\<lambda>z. f z = 0) (nhds z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
      using * by (auto simp: has_fps_expansion_def nhds_to_0' eventually_filtermap add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
    hence "eventually (\<lambda>z. f z = 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
      by (auto simp: eventually_at_filter elim: eventually_mono)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
    with assms show False
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2182
      by (auto simp: frequently_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2183
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
  from has_fps_expansion_zorder[OF * this] have eq: "zorder f z = int (subdegree (fps_expansion f z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2185
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
  moreover have "subdegree (fps_expansion f z) = 0 \<longleftrightarrow> fps_nth (fps_expansion f z) 0 \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
    using nz by (auto simp: subdegree_eq_0_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2188
  moreover have "fps_nth (fps_expansion f z) 0 = f z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2189
    by (auto simp: fps_expansion_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
  ultimately show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
    by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
lemma zorder_pos_iff':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
  assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
  shows   "zorder f z > 0 \<longleftrightarrow> f z = 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2197
  using analytic_at assms zorder_pos_iff by blast
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
lemma zorder_ge_0:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
  assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
  shows   "zorder f z \<ge> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
  have *: "(\<lambda>w. f (z + w)) has_laurent_expansion fps_to_fls (fps_expansion f z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
    using assms by (simp add: analytic_at_imp_has_fps_expansion has_laurent_expansion_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2205
  from * assms(2) have "fps_to_fls (fps_expansion f z) \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
    by (auto simp: has_laurent_expansion_def frequently_def at_to_0' eventually_filtermap add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
  with has_laurent_expansion_zorder[OF *] show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
    by (simp add: fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
lemma zorder_eq_0_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
  assumes "f analytic_on {z}" "frequently (\<lambda>w. f w \<noteq> 0) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
  shows   "zorder f z = 0 \<longleftrightarrow> f z \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2214
  using assms zorder_eq_0I zorder_pos_iff' by fastforce
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
lemma zorder_scale:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
  assumes "f analytic_on {a * z}" "eventually (\<lambda>w. f w \<noteq> 0) (at (a * z))" "a \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
  shows "zorder (\<lambda>w. f (a * w)) z = zorder f (a * z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
  from assms(1) obtain r where r: "r > 0" "f holomorphic_on ball (a * z) r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
    by (auto simp: analytic_on_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2222
  have *: "open (ball (a * z) r)" "connected (ball (a * z) r)" "a * z \<in> ball (a * z) r"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
    using r \<open>a \<noteq> 0\<close> by (auto simp: dist_norm)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2224
  from assms(2) have "eventually (\<lambda>w. f w \<noteq> 0 \<and> w \<in> ball (a * z) r - {a * z}) (at (a * z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2225
    using \<open>r > 0\<close> by (intro eventually_conj eventually_at_in_open) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
  then obtain z0 where "f z0 \<noteq> 0 \<and> z0 \<in> ball (a * z) r - {a * z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
    using eventually_happens[of _ "at (a * z)"] by force
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
  hence **: "\<exists>w\<in>ball (a * z) r. f w \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
    by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2231
  define n where "n = nat (zorder f (a * z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
  obtain r' where r':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
     "(if f (a * z) = 0 then 0 < zorder f (a * z) else zorder f (a * z) = 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
     "r' > 0" "cball (a * z) r' \<subseteq> ball (a * z) r" "zor_poly f (a * z) holomorphic_on cball (a * z) r'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
     "\<And>w. w \<in> cball (a * z) r' \<Longrightarrow>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
        f w = zor_poly f (a * z) w * (w - a * z) ^ n \<and> zor_poly f (a * z) w \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
    unfolding n_def using zorder_exist_zero[OF r(2) * **] by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
  show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
  proof (rule zorder_eqI)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
    show "open (ball z (r' / norm a))" "z \<in> ball z (r' / norm a)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
      using r \<open>r' > 0\<close> \<open>a \<noteq> 0\<close> by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
    have "(*) a ` ball z (r' / cmod a) \<subseteq> cball (a * z) r'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
    proof safe
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2245
      fix w assume "w \<in> ball z (r' / cmod a)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
      thus "a * w \<in> cball (a * z) r'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
        using dist_mult_left[of a z w] \<open>a \<noteq> 0\<close> by (auto simp: divide_simps mult_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
    thus "(\<lambda>w. a ^ n * (zor_poly f (a * z) \<circ> (\<lambda>w. a * w)) w) holomorphic_on ball z (r' / norm a)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
      using \<open>a \<noteq> 0\<close> by (intro holomorphic_on_compose_gen[OF _ r'(4)] holomorphic_intros) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
    show "a ^ n * (zor_poly f (a * z) \<circ> (\<lambda>w. a * w)) z \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
      using r' \<open>a \<noteq> 0\<close> by auto
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2253
    show "f (a * w) = a ^ n * (zor_poly f (a * z) \<circ> (*) a) w * (w - z) powi (zorder f (a * z))"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
      if "w \<in> ball z (r' / norm a)" "w \<noteq> z" for w
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
    proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2256
      have "f (a * w) = zor_poly f (a * z) (a * w) * (a * (w - z)) ^ n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
        using that r'(5)[of "a * w"] dist_mult_left[of a z w] \<open>a \<noteq> 0\<close> unfolding ring_distribs
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2258
        by (auto simp: divide_simps mult_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
      also have "\<dots> = a ^ n * zor_poly f (a * z) (a * w) * (w - z) ^ n"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2260
        by (subst power_mult_distrib) (auto simp: mult_ac)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2261
      also have "(w - z) ^ n = (w - z) powi of_nat n"
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2262
        by simp
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2263
      also have "of_nat n = zorder f (a * z)"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2264
        using r'(1) by (auto simp: n_def split: if_splits)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2265
      finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2266
        unfolding o_def n_def .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2267
    qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2268
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2269
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2270
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2271
lemma zorder_compose_aux:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2272
  assumes "isolated_singularity_at f 0" "not_essential f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2273
  assumes G: "g has_fps_expansion G" "G \<noteq> 0" "g 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2274
  assumes "eventually (\<lambda>w. f w \<noteq> 0) (at 0)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2275
  shows   "zorder (f \<circ> g) 0 = zorder f 0 * subdegree G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2276
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2277
  obtain F where F: "f has_laurent_expansion F"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2278
    using not_essential_has_laurent_expansion_0[OF assms(1,2)] by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2279
  have [simp]: "fps_nth G 0 = 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2280
   using G \<open>g 0 = 0\<close> by (simp add: has_fps_expansion_imp_0_eq_fps_nth_0)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2281
  note [simp] = \<open>G \<noteq> 0\<close> \<open>g 0 = 0\<close>
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2282
  have [simp]: "F \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2283
    using has_laurent_expansion_eventually_nonzero_iff[of f 0 F] F assms by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2284
  have FG: "(f \<circ> g) has_laurent_expansion fls_compose_fps F G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2285
    by (intro has_laurent_expansion_compose has_laurent_expansion_fps F G) auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2286
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2287
  have "zorder (f \<circ> g) 0 = fls_subdegree (fls_compose_fps F G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2288
    using has_laurent_expansion_zorder_0 [OF FG] by (auto simp: fls_compose_fps_eq_0_iff)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2289
  also have "\<dots> = fls_subdegree F * int (subdegree G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2290
    by simp
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2291
  also have "fls_subdegree F = zorder f 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2292
    using has_laurent_expansion_zorder_0 [OF F] by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2293
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2294
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2295
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2296
lemma zorder_compose:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2297
  assumes "isolated_singularity_at f (g z)" "not_essential f (g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2298
  assumes G: "(\<lambda>x. g (z + x) - g z) has_fps_expansion G" "G \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2299
  assumes "eventually (\<lambda>w. f w \<noteq> 0) (at (g z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2300
  shows   "zorder (f \<circ> g) z = zorder f (g z) * subdegree G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2301
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2302
  define f' where "f' = (\<lambda>w. f (g z + w))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2303
  define g' where "g' = (\<lambda>w. g (z + w) - g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2304
  have "zorder f (g z) = zorder f' 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2305
    by (simp add: f'_def zorder_shift' add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2306
  have "zorder (\<lambda>x. g x - g z) z = zorder g' 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2307
    by (simp add: g'_def zorder_shift' add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2308
  have "zorder (f \<circ> g) z = zorder (f' \<circ> g') 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2309
    by (simp add: zorder_shift' f'_def g'_def add_ac o_def)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2310
  also have "\<dots> = zorder f' 0 * int (subdegree G)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2311
  proof (rule zorder_compose_aux)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2312
    show "isolated_singularity_at f' 0" unfolding f'_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2313
      using assms has_laurent_expansion_isolated_0 not_essential_has_laurent_expansion by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2314
    show "not_essential f' 0" unfolding f'_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2315
      using assms has_laurent_expansion_not_essential_0 not_essential_has_laurent_expansion by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2316
  qed (use assms in \<open>auto simp: f'_def g'_def at_to_0' eventually_filtermap add_ac\<close>)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2317
  also have "zorder f' 0 = zorder f (g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2318
    by (simp add: f'_def zorder_shift' add_ac)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2319
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2320
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2321
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2322
lemma fps_to_fls_eq_fls_const_iff [simp]: "fps_to_fls F = fls_const c \<longleftrightarrow> F = fps_const c"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2323
  using fps_to_fls_eq_iff by fastforce
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2324
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2325
lemma zorder_compose':
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2326
  assumes "isolated_singularity_at f (g z)" "not_essential f (g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2327
  assumes "g analytic_on {z}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2328
  assumes "eventually (\<lambda>w. f w \<noteq> 0) (at (g z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2329
  assumes "eventually (\<lambda>w. g w \<noteq> g z) (at z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2330
  shows   "zorder (f \<circ> g) z = zorder f (g z) * zorder (\<lambda>x. g x - g z) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2331
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2332
  obtain G where G [fps_expansion_intros]: "(\<lambda>x. g (z + x)) has_fps_expansion G"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2333
    using assms analytic_at_imp_has_fps_expansion by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2334
  have G': "(\<lambda>x. g (z + x) - g z) has_fps_expansion G - fps_const (g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2335
    by (intro fps_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2336
  hence G'': "(\<lambda>x. g (z + x) - g z) has_laurent_expansion fps_to_fls (G - fps_const (g z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2337
    using has_laurent_expansion_fps by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2338
  have nz: "G - fps_const (g z) \<noteq> 0"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2339
    using has_laurent_expansion_eventually_nonzero_iff[OF G''] assms by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2340
  have "zorder (f \<circ> g) z = zorder f (g z) * subdegree (G - fps_const (g z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2341
  proof (rule zorder_compose)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2342
    show "(\<lambda>x. g (z + x) - g z) has_fps_expansion G - fps_const (g z)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2343
      by (intro fps_expansion_intros)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2344
  qed (use assms nz in auto)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2345
  also have "int (subdegree (G - fps_const (g z))) = fls_subdegree (fps_to_fls G - fls_const (g z))"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2346
    by (simp flip: fls_subdegree_fls_to_fps)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2347
  also have "\<dots> = zorder (\<lambda>x. g x - g z) z"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2348
    using has_laurent_expansion_zorder [OF G''] nz by auto
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2349
  finally show ?thesis .
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2350
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2351
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2352
lemma analytic_at_cong:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2353
  assumes "eventually (\<lambda>x. f x = g x) (nhds x)" "x = y"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2354
  shows "f analytic_on {x} \<longleftrightarrow> g analytic_on {y}"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2355
proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2356
  have "g analytic_on {x}" if "f analytic_on {x}" "eventually (\<lambda>x. f x = g x) (nhds x)" for f g
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2357
  proof -
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2358
    have "(\<lambda>y. f (x + y)) has_fps_expansion fps_expansion f x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2359
      by (rule analytic_at_imp_has_fps_expansion) fact
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2360
    also have "?this \<longleftrightarrow> (\<lambda>y. g (x + y)) has_fps_expansion fps_expansion f x"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2361
      using that by (intro has_fps_expansion_cong refl) (auto simp: nhds_to_0' eventually_filtermap)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2362
    finally show ?thesis
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2363
      by (rule has_fps_expansion_imp_analytic)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2364
  qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2365
  from this[of f g] this[of g f] show ?thesis using assms
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2366
    by (auto simp: eq_commute)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2367
qed
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2368
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2369
lemma has_laurent_expansion_sin' [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2370
  "sin has_laurent_expansion fps_to_fls (fps_sin 1)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2371
  using has_fps_expansion_sin' has_fps_expansion_to_laurent by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2372
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2373
lemma has_laurent_expansion_cos' [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2374
  "cos has_laurent_expansion fps_to_fls (fps_cos 1)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2375
  using has_fps_expansion_cos' has_fps_expansion_to_laurent by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2376
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2377
lemma has_laurent_expansion_sin [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2378
  "(\<lambda>z. sin (c * z)) has_laurent_expansion fps_to_fls (fps_sin c)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2379
  by (intro has_laurent_expansion_fps has_fps_expansion_sin)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2380
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2381
lemma has_laurent_expansion_cos [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2382
  "(\<lambda>z. cos (c * z)) has_laurent_expansion fps_to_fls (fps_cos c)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2383
  by (intro has_laurent_expansion_fps has_fps_expansion_cos)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2384
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2385
lemma has_laurent_expansion_tan' [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2386
  "tan has_laurent_expansion fps_to_fls (fps_tan 1)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2387
  using has_fps_expansion_tan' has_fps_expansion_to_laurent by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2388
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2389
lemma has_laurent_expansion_tan [laurent_expansion_intros]:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2390
  "(\<lambda>z. tan (c * z)) has_laurent_expansion fps_to_fls (fps_tan c)"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2391
  by (intro has_laurent_expansion_fps has_fps_expansion_tan)
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2392
79945
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2393
subsection \<open>More Laurent expansions\<close>
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2394
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2395
lemma has_laurent_expansion_frequently_zero_iff:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2396
  assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2397
  shows   "frequently (\<lambda>z. f z = 0) (at z) \<longleftrightarrow> F = 0"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2398
  using assms by (simp add: frequently_def has_laurent_expansion_eventually_nonzero_iff)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2399
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2400
lemma has_laurent_expansion_eventually_zero_iff:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2401
  assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2402
  shows   "eventually (\<lambda>z. f z = 0) (at z) \<longleftrightarrow> F = 0"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2403
  using assms
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2404
  by (metis has_laurent_expansion_frequently_zero_iff has_laurent_expansion_isolated 
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2405
            has_laurent_expansion_not_essential laurent_expansion_def 
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2406
            not_essential_frequently_0_imp_eventually_0 not_essential_has_laurent_expansion)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2407
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2408
lemma has_laurent_expansion_frequently_nonzero_iff:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2409
  assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2410
  shows   "frequently (\<lambda>z. f z \<noteq> 0) (at z) \<longleftrightarrow> F \<noteq> 0"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2411
  using assms by (metis has_laurent_expansion_eventually_zero_iff not_eventually)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2412
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2413
lemma has_laurent_expansion_sum_list [laurent_expansion_intros]:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2414
  assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x has_laurent_expansion F x"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2415
  shows   "(\<lambda>y. \<Sum>x\<leftarrow>xs. f x y) has_laurent_expansion (\<Sum>x\<leftarrow>xs. F x)"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2416
  using assms by (induction xs) (auto intro!: laurent_expansion_intros)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2417
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2418
lemma has_laurent_expansion_prod_list [laurent_expansion_intros]:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2419
  assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x has_laurent_expansion F x"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2420
  shows   "(\<lambda>y. \<Prod>x\<leftarrow>xs. f x y) has_laurent_expansion (\<Prod>x\<leftarrow>xs. F x)"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2421
  using assms by (induction xs) (auto intro!: laurent_expansion_intros)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2422
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2423
lemma has_laurent_expansion_sum_mset [laurent_expansion_intros]:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2424
  assumes "\<And>x. x \<in># I \<Longrightarrow> f x has_laurent_expansion F x"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2425
  shows   "(\<lambda>y. \<Sum>x\<in>#I. f x y) has_laurent_expansion (\<Sum>x\<in>#I. F x)"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2426
  using assms by (induction I) (auto intro!: laurent_expansion_intros)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2427
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2428
lemma has_laurent_expansion_prod_mset [laurent_expansion_intros]:
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2429
  assumes "\<And>x. x \<in># I \<Longrightarrow> f x has_laurent_expansion F x"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2430
  shows   "(\<lambda>y. \<Prod>x\<in>#I. f x y) has_laurent_expansion (\<Prod>x\<in>#I. F x)"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2431
  using assms by (induction I) (auto intro!: laurent_expansion_intros)
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78751
diff changeset
  2432
82530
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2433
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2434
subsection \<open>Formal convergence versus analytic convergence\<close>
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2435
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2436
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2437
text \<open>
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2438
  The convergence of a sequence of formal power series and the convergence of the functions
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2439
  in the complex plane do not imply each other:
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2440
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2441
  \<^item> If we have the sequence of constant power series $(1/n)_{n\geq 0}$, this clearly converges 
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2442
    to the zero function analytically, but as a series of formal power series it is divergent 
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2443
    (since the 0-th coefficient never stabilises).
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2444
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2445
  \<^item> Conversely, the sequence of series $(n! x^n)_{n\geq 0}$ converges formally to $0$,
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2446
    but the corresponding sequence of functions diverges for every $x \neq 0$.
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2447
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2448
  However, if the sequence of series converges to some limit series $h$ and the corresponding
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2449
  series of functions converges uniformly to some limit function $g(x)$, then $h$ is also a
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2450
  series expansion of $g(x)$, i.e.\ in that case, formal and analytic convergence agree.
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2451
\<close>
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2452
proposition uniform_limit_imp_fps_expansion_eq:
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2453
  fixes f :: "'a \<Rightarrow> complex fps"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2454
  assumes lim1: "(f \<longlongrightarrow> h) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2455
  assumes lim2: "uniform_limit A (\<lambda>x z. f' x z) g' F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2456
  assumes expansions: "eventually (\<lambda>x. f' x has_fps_expansion f x) F" "g' has_fps_expansion g"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2457
  assumes holo: "eventually (\<lambda>x. f' x holomorphic_on A) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2458
  assumes A: "open A" "0 \<in> A"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2459
  assumes nontriv [simp]: "F \<noteq> bot"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2460
  shows "g = h"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2461
proof (rule fps_ext)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2462
  fix n :: nat
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2463
  have "eventually (\<lambda>x. fps_nth (f x) n = fps_nth h n) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2464
    using lim1 unfolding tendsto_fps_iff by blast
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2465
  hence "eventually (\<lambda>x. (deriv ^^ n) (f' x) 0 / fact n = fps_nth h n) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2466
    using expansions(1)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2467
  proof eventually_elim
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2468
    case (elim x)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2469
    have "fps_nth (f x) n = (deriv ^^ n) (f' x) 0 / fact n"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2470
      by (rule fps_nth_fps_expansion) (use elim in auto)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2471
    with elim show ?case
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2472
      by simp
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2473
  qed
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2474
  hence "((\<lambda>x. (deriv ^^ n) (f' x) 0 / fact n) \<longlongrightarrow> fps_nth h n) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2475
    by (simp add: tendsto_eventually)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2476
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2477
  moreover have "((\<lambda>x. (deriv ^^ n) (f' x) 0) \<longlongrightarrow> (deriv ^^ n) g' 0) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2478
    using lim2
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2479
  proof (rule higher_deriv_complex_uniform_limit)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2480
    show "eventually (\<lambda>x. f' x holomorphic_on A) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2481
      using holo by eventually_elim auto
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2482
  qed (use A in auto)
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2483
  hence "((\<lambda>x. (deriv ^^ n) (f' x) 0 / fact n) \<longlongrightarrow> (deriv ^^ n) g' 0 / fact n) F"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2484
    by (intro tendsto_divide) auto
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2485
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2486
  ultimately have "fps_nth h n = (deriv ^^ n) g' 0 / fact n"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2487
    using tendsto_unique[OF nontriv] by blast
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2488
  also have "\<dots> = fps_nth g n"
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2489
    by (rule fps_nth_fps_expansion [symmetric]) fact
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2490
  finally show "fps_nth g n = fps_nth h n" ..
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2491
qed
904589510439 some facts about power series
Manuel Eberl <manuel@pruvisto.org>
parents: 82529
diff changeset
  2492
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2493
end