src/HOL/Analysis/Complex_Analysis_Basics.thy
author paulson <lp15@cam.ac.uk>
Sat, 12 Aug 2023 10:09:29 +0100
changeset 78516 56a408fa2440
parent 77226 69956724ad4f
child 79857 819c28a7280f
permissions -rw-r--r--
substantial tidy-up, shortening many proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(*  Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
    Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
*)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
     5
section \<open>Complex Analysis Basics\<close>
71167
b4d409c65a76 Rearrangement of material in Complex_Analysis_Basics, which contained much that had nothing to do with complex analysis.
paulson <lp15@cam.ac.uk>
parents: 71030
diff changeset
     6
text \<open>Definitions of analytic and holomorphic functions, limit theorems, complex differentiation\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
theory Complex_Analysis_Basics
77200
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
     9
  imports Derivative "HOL-Library.Nonpos_Ints" Uncountable_Sets
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
    12
subsection\<^marker>\<open>tag unimportant\<close>\<open>General lemmas\<close>
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
    13
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
    14
lemma nonneg_Reals_cmod_eq_Re: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> norm z = Re z"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
    15
  by (simp add: complex_nonneg_Reals_iff cmod_eq_Re)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    16
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
    17
lemma fact_cancel:
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
    18
  fixes c :: "'a::real_field"
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
    19
  shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
    20
  using of_nat_neq_0 by force
56889
48a745e1bde7 avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
hoelzl
parents: 56479
diff changeset
    21
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    22
lemma vector_derivative_cnj_within:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    23
  assumes "at x within A \<noteq> bot" and "f differentiable at x within A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    24
  shows   "vector_derivative (\<lambda>z. cnj (f z)) (at x within A) = 
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    25
             cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D")
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    26
proof -
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    27
  let ?D = "vector_derivative f (at x within A)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    28
  from assms have "(f has_vector_derivative ?D) (at x within A)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    29
    by (subst (asm) vector_derivative_works)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    30
  hence "((\<lambda>x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    31
    by (rule has_vector_derivative_cnj)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    32
  thus ?thesis using assms by (auto dest: vector_derivative_within)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    33
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    34
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    35
lemma vector_derivative_cnj:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    36
  assumes "f differentiable at x"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    37
  shows   "vector_derivative (\<lambda>z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    38
  using assms by (intro vector_derivative_cnj_within) auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
    39
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    40
lemma
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    41
  shows open_halfspace_Re_lt: "open {z. Re(z) < b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    42
    and open_halfspace_Re_gt: "open {z. Re(z) > b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    43
    and closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    44
    and closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    45
    and closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    46
    and open_halfspace_Im_lt: "open {z. Im(z) < b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    47
    and open_halfspace_Im_gt: "open {z. Im(z) > b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    48
    and closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    49
    and closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    50
    and closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
63332
f164526d8727 move open_Collect_eq/less to HOL
hoelzl
parents: 63092
diff changeset
    51
  by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re
f164526d8727 move open_Collect_eq/less to HOL
hoelzl
parents: 63092
diff changeset
    52
            continuous_on_Im continuous_on_id continuous_on_const)+
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    53
77200
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    54
lemma uncountable_halfspace_Im_gt: "uncountable {z. Im z > c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    55
proof -
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    56
  obtain r where r: "r > 0" "ball ((c + 1) *\<^sub>R \<i>) r \<subseteq> {z. Im z > c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    57
    using open_halfspace_Im_gt[of c] unfolding open_contains_ball by force
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    58
  then show ?thesis
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    59
    using countable_subset uncountable_ball by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    60
qed
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    61
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    62
lemma uncountable_halfspace_Im_lt: "uncountable {z. Im z < c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    63
proof -
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    64
  obtain r where r: "r > 0" "ball ((c - 1) *\<^sub>R \<i>) r \<subseteq> {z. Im z < c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    65
    using open_halfspace_Im_lt[of c] unfolding open_contains_ball by force
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    66
  then show ?thesis
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    67
    using countable_subset uncountable_ball by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    68
qed
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    69
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    70
lemma uncountable_halfspace_Re_gt: "uncountable {z. Re z > c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    71
proof -
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    72
  obtain r where r: "r > 0" "ball (of_real(c + 1)) r \<subseteq> {z. Re z > c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    73
    using open_halfspace_Re_gt[of c] unfolding open_contains_ball by force
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    74
  then show ?thesis
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    75
    using countable_subset uncountable_ball by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    76
qed
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    77
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    78
lemma uncountable_halfspace_Re_lt: "uncountable {z. Re z < c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    79
proof -
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    80
  obtain r where r: "r > 0" "ball (of_real(c - 1)) r \<subseteq> {z. Re z < c}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    81
    using open_halfspace_Re_lt[of c] unfolding open_contains_ball by force
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    82
  then show ?thesis
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    83
    using countable_subset uncountable_ball by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    84
qed
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    85
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    86
lemma connected_halfspace_Im_gt [intro]: "connected {z. c < Im z}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    87
  by (intro convex_connected convex_halfspace_Im_gt)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    88
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    89
lemma connected_halfspace_Im_lt [intro]: "connected {z. c > Im z}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    90
  by (intro convex_connected convex_halfspace_Im_lt)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    91
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    92
lemma connected_halfspace_Re_gt [intro]: "connected {z. c < Re z}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    93
  by (intro convex_connected convex_halfspace_Re_gt)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    94
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    95
lemma connected_halfspace_Re_lt [intro]: "connected {z. c > Re z}"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    96
  by (intro convex_connected convex_halfspace_Re_lt)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
    97
  
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60585
diff changeset
    98
lemma closed_complex_Reals: "closed (\<real> :: complex set)"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
    99
proof -
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60585
diff changeset
   100
  have "(\<real> :: complex set) = {z. Im z = 0}"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   101
    by (auto simp: complex_is_Real_iff)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   102
  then show ?thesis
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   103
    by (metis closed_halfspace_Im_eq)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   104
qed
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   105
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   106
lemma closed_Real_halfspace_Re_le: "closed (\<real> \<inter> {w. Re w \<le> x})"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   107
  by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   108
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   109
lemma closed_nonpos_Reals_complex [simp]: "closed (\<real>\<^sub>\<le>\<^sub>0 :: complex set)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   110
proof -
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   111
  have "\<real>\<^sub>\<le>\<^sub>0 = \<real> \<inter> {z. Re(z) \<le> 0}"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   112
    using complex_nonpos_Reals_iff complex_is_Real_iff by auto
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   113
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   114
    by (metis closed_Real_halfspace_Re_le)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   115
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   116
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   117
lemma closed_Real_halfspace_Re_ge: "closed (\<real> \<inter> {w. x \<le> Re(w)})"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   118
  using closed_halfspace_Re_ge
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   119
  by (simp add: closed_Int closed_complex_Reals)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   120
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 69064
diff changeset
   121
lemma closed_nonneg_Reals_complex [simp]: "closed (\<real>\<^sub>\<ge>\<^sub>0 :: complex set)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   122
proof -
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   123
  have "\<real>\<^sub>\<ge>\<^sub>0 = \<real> \<inter> {z. Re(z) \<ge> 0}"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   124
    using complex_nonneg_Reals_iff complex_is_Real_iff by auto
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   125
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   126
    by (metis closed_Real_halfspace_Re_ge)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   127
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   128
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   129
lemma closed_real_abs_le: "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   130
proof -
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   131
  have "{w \<in> \<real>. \<bar>Re w\<bar> \<le> r} = (\<real> \<inter> {w. Re w \<le> r}) \<inter> (\<real> \<inter> {w. Re w \<ge> -r})"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   132
    by auto
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   133
  then show "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   134
    by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   135
qed
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   136
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   137
lemma real_lim:
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   138
  fixes l::complex
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
   139
  assumes "(f \<longlongrightarrow> l) F" and "\<not> trivial_limit F" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   140
  shows  "l \<in> \<real>"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   141
  using Lim_in_closed_set[OF closed_complex_Reals] assms
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   142
  by (smt (verit) eventually_mono)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   143
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   144
lemma real_lim_sequentially:
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   145
  fixes l::complex
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   146
  shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   147
  by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   148
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   149
lemma real_series:
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   150
  fixes l::complex
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   151
  shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   152
  unfolding sums_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   153
  by (metis real_lim_sequentially sum_in_Reals)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   154
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   155
lemma Lim_null_comparison_Re:
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   156
  assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   157
  using Lim_null_comparison assms tendsto_Re by fastforce
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   159
subsection\<open>Holomorphic functions\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   161
definition\<^marker>\<open>tag important\<close> holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   162
           (infixl "(holomorphic'_on)" 50)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   163
  where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   164
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   165
named_theorems\<^marker>\<open>tag important\<close> holomorphic_intros "structural introduction rules for holomorphic_on"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   166
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   167
lemma holomorphic_onI [intro?]: "(\<And>x. x \<in> s \<Longrightarrow> f field_differentiable (at x within s)) \<Longrightarrow> f holomorphic_on s"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   168
  by (simp add: holomorphic_on_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   169
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   170
lemma holomorphic_onD [dest?]: "\<lbrakk>f holomorphic_on s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x within s)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   171
  by (simp add: holomorphic_on_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   172
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   173
lemma holomorphic_on_imp_differentiable_on:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   174
    "f holomorphic_on s \<Longrightarrow> f differentiable_on s"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   175
  unfolding holomorphic_on_def differentiable_on_def
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   176
  by (simp add: field_differentiable_imp_differentiable)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   177
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   178
lemma holomorphic_on_imp_differentiable_at:
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   179
   "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   180
using at_within_open holomorphic_on_def by fastforce
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
   181
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   182
lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   183
  by (simp add: holomorphic_on_def)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   184
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   185
lemma holomorphic_on_open:
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   186
    "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   187
  by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s])
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   188
74513
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   189
lemma holomorphic_on_UN_open:
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   190
  assumes "\<And>n. n \<in> I \<Longrightarrow> f holomorphic_on A n" "\<And>n. n \<in> I \<Longrightarrow> open (A n)"
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   191
  shows   "f holomorphic_on (\<Union>n\<in>I. A n)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   192
  by (metis UN_E assms holomorphic_on_open open_UN)
74513
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   193
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   194
lemma holomorphic_on_imp_continuous_on:
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   195
    "f holomorphic_on s \<Longrightarrow> continuous_on s f"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   196
  using differentiable_imp_continuous_on holomorphic_on_imp_differentiable_on by blast
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   197
73885
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   198
lemma holomorphic_closedin_preimage_constant:
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   199
  assumes "f holomorphic_on D" 
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   200
  shows "closedin (top_of_set D) {z\<in>D. f z = a}"
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   201
  by (simp add: assms continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on)
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   202
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   203
lemma holomorphic_closed_preimage_constant:
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   204
  assumes "f holomorphic_on UNIV" 
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   205
  shows "closed {z. f z = a}"
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   206
  using holomorphic_closedin_preimage_constant [OF assms] by simp
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   207
62540
f2fc5485e3b0 Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
   208
lemma holomorphic_on_subset [elim]:
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   209
    "f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   210
  unfolding holomorphic_on_def
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   211
  by (metis field_differentiable_within_subset subsetD)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   212
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   213
lemma holomorphic_transform: "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   214
  by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   215
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   216
lemma holomorphic_cong: "s = t ==> (\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on t"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   217
  by (metis holomorphic_transform)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   218
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68721
diff changeset
   219
lemma holomorphic_on_linear [simp, holomorphic_intros]: "((*) c) holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   220
  unfolding holomorphic_on_def by (metis field_differentiable_linear)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   221
62217
527488dc8b90 Reorganised a huge proof
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   222
lemma holomorphic_on_const [simp, holomorphic_intros]: "(\<lambda>z. c) holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   223
  unfolding holomorphic_on_def by (metis field_differentiable_const)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   224
62217
527488dc8b90 Reorganised a huge proof
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   225
lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\<lambda>x. x) holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   226
  unfolding holomorphic_on_def by (metis field_differentiable_ident)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   227
62217
527488dc8b90 Reorganised a huge proof
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
   228
lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   229
  unfolding id_def by (rule holomorphic_on_ident)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   230
74513
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   231
lemma constant_on_imp_holomorphic_on:
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   232
  assumes "f constant_on A"
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   233
  shows   "f holomorphic_on A"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   234
  by (metis assms constant_on_def holomorphic_on_const holomorphic_transform)
74513
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
   235
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   236
lemma holomorphic_on_compose:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   237
  "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g \<circ> f) holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   238
  using field_differentiable_compose_within[of f _ s g]
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   239
  by (auto simp: holomorphic_on_def)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   240
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   241
lemma holomorphic_on_compose_gen:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   242
  "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g \<circ> f) holomorphic_on s"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   243
  by (metis holomorphic_on_compose holomorphic_on_subset)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   244
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   245
lemma holomorphic_on_balls_imp_entire:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   246
  assumes "\<not>bdd_above A" "\<And>r. r \<in> A \<Longrightarrow> f holomorphic_on ball c r"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   247
  shows   "f holomorphic_on B"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   248
proof (rule holomorphic_on_subset)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   249
  show "f holomorphic_on UNIV" unfolding holomorphic_on_def
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   250
  proof
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   251
    fix z :: complex
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   252
    from \<open>\<not>bdd_above A\<close> obtain r where r: "r \<in> A" "r > norm (z - c)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   253
      by (meson bdd_aboveI not_le)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   254
    with assms(2) have "f holomorphic_on ball c r" by blast
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   255
    moreover from r have "z \<in> ball c r" by (auto simp: dist_norm norm_minus_commute)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   256
    ultimately show "f field_differentiable at z"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   257
      by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"])
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   258
  qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   259
qed auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   260
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   261
lemma holomorphic_on_balls_imp_entire':
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   262
  assumes "\<And>r. r > 0 \<Longrightarrow> f holomorphic_on ball c r"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   263
  shows   "f holomorphic_on B"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   264
proof (rule holomorphic_on_balls_imp_entire)  
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   265
  show "\<not>bdd_above {(0::real)<..}" unfolding bdd_above_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   266
    by (meson greaterThan_iff gt_ex less_le_not_le order_le_less_trans)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   267
qed (use assms in auto)
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68296
diff changeset
   268
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   269
lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on A \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   270
  by (metis field_differentiable_minus holomorphic_on_def)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   271
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   272
lemma holomorphic_on_add [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   273
  "\<lbrakk>f holomorphic_on A; g holomorphic_on A\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   274
  unfolding holomorphic_on_def by (metis field_differentiable_add)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   275
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   276
lemma holomorphic_on_diff [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   277
  "\<lbrakk>f holomorphic_on A; g holomorphic_on A\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   278
  unfolding holomorphic_on_def by (metis field_differentiable_diff)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   279
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   280
lemma holomorphic_on_mult [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   281
  "\<lbrakk>f holomorphic_on A; g holomorphic_on A\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   282
  unfolding holomorphic_on_def by (metis field_differentiable_mult)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   283
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   284
lemma holomorphic_on_inverse [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   285
  "\<lbrakk>f holomorphic_on A; \<And>z. z \<in> A \<Longrightarrow> f z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. inverse (f z)) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   286
  unfolding holomorphic_on_def by (metis field_differentiable_inverse)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   287
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   288
lemma holomorphic_on_divide [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   289
  "\<lbrakk>f holomorphic_on A; g holomorphic_on A; \<And>z. z \<in> A \<Longrightarrow> g z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. f z / g z) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   290
  unfolding holomorphic_on_def by (metis field_differentiable_divide)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   291
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
   292
lemma holomorphic_on_power [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   293
  "f holomorphic_on A \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   294
  unfolding holomorphic_on_def by (metis field_differentiable_power)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   295
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   296
lemma holomorphic_on_power_int [holomorphic_intros]:
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   297
  assumes nz: "n \<ge> 0 \<or> (\<forall>x\<in>A. f x \<noteq> 0)" and f: "f holomorphic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   298
  shows   "(\<lambda>x. f x powi n) holomorphic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   299
proof (cases "n \<ge> 0")
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   300
  case True
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   301
  have "(\<lambda>x. f x ^ nat n) holomorphic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   302
    by (simp add: f holomorphic_on_power)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   303
  with True show ?thesis
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   304
    by (simp add: power_int_def)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   305
next
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   306
  case False
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   307
  hence "(\<lambda>x. inverse (f x ^ nat (-n))) holomorphic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   308
    using nz by (auto intro!: holomorphic_intros f)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   309
  with False show ?thesis
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   310
    by (simp add: power_int_def power_inverse)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   311
qed
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   312
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   313
lemma holomorphic_on_sum [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   314
  "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on A) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) holomorphic_on A"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   315
  unfolding holomorphic_on_def by (metis field_differentiable_sum)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   316
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
   317
lemma holomorphic_on_prod [holomorphic_intros]:
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   318
  "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on A) \<Longrightarrow> (\<lambda>x. prod (\<lambda>i. f i x) I) holomorphic_on A"
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
   319
  by (induction I rule: infinite_finite_induct) (auto intro: holomorphic_intros)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
   320
66486
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   321
lemma holomorphic_pochhammer [holomorphic_intros]:
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   322
  "f holomorphic_on A \<Longrightarrow> (\<lambda>s. pochhammer (f s) n) holomorphic_on A"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   323
  by (induction n) (auto intro!: holomorphic_intros simp: pochhammer_Suc)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   324
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   325
lemma holomorphic_on_scaleR [holomorphic_intros]:
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   326
  "f holomorphic_on A \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) holomorphic_on A"
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   327
  by (auto simp: scaleR_conv_of_real intro!: holomorphic_intros)
ffaaa83543b2 Lemmas about analysis and permutations
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   328
67167
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   329
lemma holomorphic_on_Un [holomorphic_intros]:
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   330
  assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   331
  shows   "f holomorphic_on (A \<union> B)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   332
  by (metis Un_iff assms holomorphic_on_open open_Un)
67167
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   333
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   334
lemma holomorphic_on_If_Un [holomorphic_intros]:
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   335
  assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   336
  assumes "\<And>z. z \<in> A \<Longrightarrow> z \<in> B \<Longrightarrow> f z = g z"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   337
  shows   "(\<lambda>z. if z \<in> A then f z else g z) holomorphic_on (A \<union> B)" (is "?h holomorphic_on _")
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   338
proof (intro holomorphic_on_Un)
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   339
  note \<open>f holomorphic_on A\<close>
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   340
  also have "f holomorphic_on A \<longleftrightarrow> ?h holomorphic_on A"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   341
    by (intro holomorphic_cong) auto
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   342
  finally show \<dots> .
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   343
next
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   344
  note \<open>g holomorphic_on B\<close>
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   345
  also have "g holomorphic_on B \<longleftrightarrow> ?h holomorphic_on B"
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   346
    using assms by (intro holomorphic_cong) auto
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   347
  finally show \<dots> .
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   348
qed (use assms in auto)
67167
88d1c9d86f48 Moved analysis material from AFP
Manuel Eberl <eberlm@in.tum.de>
parents: 67135
diff changeset
   349
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   350
lemma holomorphic_derivI:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   351
     "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk> \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   352
  by (metis DERIV_deriv_iff_field_differentiable at_within_open  holomorphic_on_def has_field_derivative_at_within)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   353
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   354
lemma complex_derivative_transform_within_open:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   355
  "\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk>
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   356
   \<Longrightarrow> deriv f z = deriv g z"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   357
  by (smt (verit) DERIV_imp_deriv has_field_derivative_transform_within_open holomorphic_on_open)
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   358
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   359
lemma holomorphic_on_compose_cnj_cnj:
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   360
  assumes "f holomorphic_on cnj ` A" "open A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   361
  shows   "cnj \<circ> f \<circ> cnj holomorphic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   362
proof -
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   363
  have [simp]: "open (cnj ` A)"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   364
    unfolding image_cnj_conv_vimage_cnj using assms by (intro open_vimage) auto
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   365
  show ?thesis
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   366
    using assms unfolding holomorphic_on_def
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   367
    by (auto intro!: field_differentiable_cnj_cnj simp: at_within_open_NO_MATCH)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   368
qed
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   369
  
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   370
lemma holomorphic_nonconstant:
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   371
  assumes holf: "f holomorphic_on S" and "open S" "\<xi> \<in> S" "deriv f \<xi> \<noteq> 0"
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   372
    shows "\<not> f constant_on S"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   373
  by (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S])
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   374
    (use assms in \<open>auto simp: holomorphic_derivI\<close>)
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62408
diff changeset
   375
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   376
subsection\<open>Analyticity on a set\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   378
definition\<^marker>\<open>tag important\<close> analytic_on (infixl "(analytic'_on)" 50)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   379
  where "f analytic_on S \<equiv> \<forall>x \<in> S. \<exists>\<epsilon>. 0 < \<epsilon> \<and> f holomorphic_on (ball x \<epsilon>)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   381
named_theorems\<^marker>\<open>tag important\<close> analytic_intros "introduction rules for proving analyticity"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   382
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   383
lemma analytic_imp_holomorphic: "f analytic_on S \<Longrightarrow> f holomorphic_on S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   384
  unfolding analytic_on_def holomorphic_on_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   385
  using centre_in_ball field_differentiable_at_within field_differentiable_within_open by blast
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   387
lemma analytic_on_open: "open S \<Longrightarrow> f analytic_on S \<longleftrightarrow> f holomorphic_on S"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   388
  by (meson analytic_imp_holomorphic analytic_on_def holomorphic_on_subset openE)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
lemma analytic_on_imp_differentiable_at:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   391
  "f analytic_on S \<Longrightarrow> x \<in> S \<Longrightarrow> f field_differentiable (at x)"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   392
  using analytic_on_def holomorphic_on_imp_differentiable_at by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   393
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   394
lemma analytic_at_imp_isCont:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   395
  assumes "f analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   396
  shows   "isCont f z"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   397
  by (meson analytic_on_imp_differentiable_at assms field_differentiable_imp_continuous_at insertCI)
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   398
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   399
lemma analytic_at_neq_imp_eventually_neq:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   400
  assumes "f analytic_on {x}" "f x \<noteq> c"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   401
  shows   "eventually (\<lambda>y. f y \<noteq> c) (at x)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   402
  using analytic_at_imp_isCont assms isContD tendsto_imp_eventually_ne by blast
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   404
lemma analytic_on_subset: "f analytic_on S \<Longrightarrow> T \<subseteq> S \<Longrightarrow> f analytic_on T"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
  by (auto simp: analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   407
lemma analytic_on_Un: "f analytic_on (S \<union> T) \<longleftrightarrow> f analytic_on S \<and> f analytic_on T"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
  by (auto simp: analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   410
lemma analytic_on_Union: "f analytic_on (\<Union>\<T>) \<longleftrightarrow> (\<forall>T \<in> \<T>. f analytic_on T)"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   411
  by (auto simp: analytic_on_def)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   412
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   413
lemma analytic_on_UN: "f analytic_on (\<Union>i\<in>I. S i) \<longleftrightarrow> (\<forall>i\<in>I. f analytic_on (S i))"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
  by (auto simp: analytic_on_def)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   415
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
lemma analytic_on_holomorphic:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   417
  "f analytic_on S \<longleftrightarrow> (\<exists>T. open T \<and> S \<subseteq> T \<and> f holomorphic_on T)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
  (is "?lhs = ?rhs")
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
proof -
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   420
  have "?lhs \<longleftrightarrow> (\<exists>T. open T \<and> S \<subseteq> T \<and> f analytic_on T)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
  proof safe
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   422
    assume "f analytic_on S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   423
    then have "\<forall>x \<in> \<Union>{U. open U \<and> f analytic_on U}. \<exists>\<epsilon>>0. f holomorphic_on ball x \<epsilon>"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   424
      using analytic_on_def by force
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   425
    moreover have "S \<subseteq> \<Union>{U. open U \<and> f analytic_on U}"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   426
      using \<open>f analytic_on S\<close>
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   427
      by (smt (verit, best) open_ball Union_iff analytic_on_def analytic_on_open centre_in_ball mem_Collect_eq subsetI)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   428
    ultimately show "\<exists>T. open T \<and> S \<subseteq> T \<and> f analytic_on T"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   429
      unfolding analytic_on_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   430
      by (metis (mono_tags, lifting) mem_Collect_eq open_Union)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
  next
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   432
    fix T
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   433
    assume "open T" "S \<subseteq> T" "f analytic_on T"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   434
    then show "f analytic_on S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
        by (metis analytic_on_subset)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
  qed
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   437
  also have "\<dots> \<longleftrightarrow> ?rhs"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
    by (auto simp: analytic_on_open)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
  finally show ?thesis .
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68721
diff changeset
   442
lemma analytic_on_linear [analytic_intros,simp]: "((*) c) analytic_on S"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   443
  by (auto simp add: analytic_on_holomorphic)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   445
lemma analytic_on_const [analytic_intros,simp]: "(\<lambda>z. c) analytic_on S"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   446
  by (metis analytic_on_def holomorphic_on_const zero_less_one)
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   447
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   448
lemma analytic_on_ident [analytic_intros,simp]: "(\<lambda>x. x) analytic_on S"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   449
  by (simp add: analytic_on_def gt_ex)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   451
lemma analytic_on_id [analytic_intros]: "id analytic_on S"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   452
  unfolding id_def by (rule analytic_on_ident)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
lemma analytic_on_compose:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   455
  assumes f: "f analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   456
      and g: "g analytic_on (f ` S)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   457
    shows "(g \<circ> f) analytic_on S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
unfolding analytic_on_def
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
proof (intro ballI)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
  fix x
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   461
  assume x: "x \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
    by (metis analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   465
    by (metis analytic_on_def g image_eqI x)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
  have "isCont f x"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   467
    by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
  with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
     by (auto simp: continuous_at_ball)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   470
  have "g \<circ> f holomorphic_on ball x (min d e)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   471
    by (meson fd fh gh holomorphic_on_compose_gen holomorphic_on_subset image_mono min.cobounded1 min.cobounded2 subset_ball)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
  then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   473
    by (metis d e min_less_iff_conj)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
lemma analytic_on_compose_gen:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   477
  "f analytic_on S \<Longrightarrow> g analytic_on T \<Longrightarrow> (\<And>z. z \<in> S \<Longrightarrow> f z \<in> T)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   478
             \<Longrightarrow> g \<circ> f analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   479
  by (metis analytic_on_compose analytic_on_subset image_subset_iff)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   481
lemma analytic_on_neg [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   482
  "f analytic_on S \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   483
  by (metis analytic_on_holomorphic holomorphic_on_minus)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   485
lemma analytic_on_add [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   486
  assumes f: "f analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   487
      and g: "g analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   488
    shows "(\<lambda>z. f z + g z) analytic_on S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
unfolding analytic_on_def
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
proof (intro ballI)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
  fix z
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   492
  assume z: "z \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
    by (metis analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   496
    by (metis analytic_on_def g z)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   497
  have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   498
    by (metis fh gh holomorphic_on_add holomorphic_on_subset linorder_linear min_def subset_ball)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
  then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
    by (metis e e' min_less_iff_conj)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   503
lemma analytic_on_mult [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   504
  assumes f: "f analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   505
      and g: "g analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   506
    shows "(\<lambda>z. f z * g z) analytic_on S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
unfolding analytic_on_def
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
proof (intro ballI)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
  fix z
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   510
  assume z: "z \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
    by (metis analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
  obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   514
    by (metis analytic_on_def g z)
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   515
  have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   516
    by (metis fh gh holomorphic_on_mult holomorphic_on_subset min.absorb_iff2 min_def subset_ball)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
  then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
    by (metis e e' min_less_iff_conj)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   521
lemma analytic_on_diff [analytic_intros]:
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   522
  assumes f: "f analytic_on S" and g: "g analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   523
  shows "(\<lambda>z. f z - g z) analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   524
proof -
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   525
  have "(\<lambda>z. - g z) analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   526
    by (simp add: analytic_on_neg g)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   527
  then have "(\<lambda>z. f z + - g z) analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   528
    using analytic_on_add f by blast
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   529
  then show ?thesis
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   530
    by fastforce
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   531
qed
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   532
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   533
lemma analytic_on_inverse [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   534
  assumes f: "f analytic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   535
      and nz: "(\<And>z. z \<in> S \<Longrightarrow> f z \<noteq> 0)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   536
    shows "(\<lambda>z. inverse (f z)) analytic_on S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
unfolding analytic_on_def
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
proof (intro ballI)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
  fix z
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   540
  assume z: "z \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
  then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    by (metis analytic_on_def)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
  have "continuous_on (ball z e) f"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
    by (metis fh holomorphic_on_imp_continuous_on)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   545
  then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66486
diff changeset
   546
    by (metis open_ball centre_in_ball continuous_on_open_avoid e z nz)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   547
  have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   548
    using fh holomorphic_on_inverse holomorphic_on_open nz' by fastforce
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
  then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
    by (metis e e' min_less_iff_conj)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   553
lemma analytic_on_divide [analytic_intros]:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   554
  assumes f: "f analytic_on S" and g: "g analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   555
    and nz: "(\<And>z. z \<in> S \<Longrightarrow> g z \<noteq> 0)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   556
  shows "(\<lambda>z. f z / g z) analytic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   557
  unfolding divide_inverse by (metis analytic_on_inverse analytic_on_mult f g nz)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   558
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   559
lemma analytic_on_power [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   560
  "f analytic_on S \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   561
  by (induct n) (auto simp: analytic_on_mult)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   562
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   563
lemma analytic_on_power_int [analytic_intros]:
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   564
  assumes nz: "n \<ge> 0 \<or> (\<forall>x\<in>A. f x \<noteq> 0)" and f: "f analytic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   565
  shows   "(\<lambda>x. f x powi n) analytic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   566
proof (cases "n \<ge> 0")
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   567
  case True
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   568
  have "(\<lambda>x. f x ^ nat n) analytic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   569
    using analytic_on_power f by blast
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   570
  with True show ?thesis
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   571
    by (simp add: power_int_def)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   572
next
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   573
  case False
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   574
  hence "(\<lambda>x. inverse (f x ^ nat (-n))) analytic_on A"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   575
    using nz by (auto intro!: analytic_intros f)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   576
  with False show ?thesis
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   577
    by (simp add: power_int_def power_inverse)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   578
qed
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 75168
diff changeset
   579
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 64394
diff changeset
   580
lemma analytic_on_sum [analytic_intros]:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   581
  "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on S) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) analytic_on S"
71633
07bec530f02e cleaned proofs
nipkow
parents: 71189
diff changeset
   582
  by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_add)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
75168
ff60b4acd6dd Added some theorems (from Wetzel)
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   584
lemma analytic_on_prod [analytic_intros]:
ff60b4acd6dd Added some theorems (from Wetzel)
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   585
  "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on S) \<Longrightarrow> (\<lambda>x. prod (\<lambda>i. f i x) I) analytic_on S"
ff60b4acd6dd Added some theorems (from Wetzel)
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   586
  by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_mult)
ff60b4acd6dd Added some theorems (from Wetzel)
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   587
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   588
lemma deriv_left_inverse:
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   589
  assumes "f holomorphic_on S" and "g holomorphic_on T"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   590
      and "open S" and "open T"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   591
      and "f ` S \<subseteq> T"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   592
      and [simp]: "\<And>z. z \<in> S \<Longrightarrow> g (f z) = z"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   593
      and "w \<in> S"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   594
    shows "deriv f w * deriv g (f w) = 1"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   595
proof -
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   596
  have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w"
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   597
    by (simp add: algebra_simps)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   598
  also have "\<dots> = deriv (g \<circ> f) w"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   599
    using assms
71189
954ee5acaae0 Split off new HOL-Complex_Analysis session from HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 71167
diff changeset
   600
    by (metis analytic_on_imp_differentiable_at analytic_on_open deriv_chain image_subset_iff)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   601
  also have "\<dots> = deriv id w"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   602
  proof (rule complex_derivative_transform_within_open [where s=S])
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   603
    show "g \<circ> f holomorphic_on S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   604
      by (rule assms holomorphic_on_compose_gen holomorphic_intros)+
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   605
  qed (use assms in auto)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   606
  also have "\<dots> = 1"
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   607
    by simp
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   608
  finally show ?thesis .
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   609
qed
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   610
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   611
subsection\<^marker>\<open>tag unimportant\<close>\<open>Analyticity at a point\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   613
lemma analytic_at_ball:
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   614
  "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   615
  by (metis analytic_on_def singleton_iff)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   616
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
lemma analytic_at:
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   618
  "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   619
  by (metis analytic_on_holomorphic empty_subsetI insert_subset)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   620
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   621
lemma holomorphic_on_imp_analytic_at:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   622
  assumes "f holomorphic_on A" "open A" "z \<in> A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   623
  shows   "f analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   624
  using assms by (meson analytic_at)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   626
lemma analytic_on_analytic_at:
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   627
  "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   628
  by (metis analytic_at_ball analytic_on_def)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   630
lemma analytic_at_two:
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   631
  "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   632
   (\<exists>S. open S \<and> z \<in> S \<and> f holomorphic_on S \<and> g holomorphic_on S)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   633
  (is "?lhs = ?rhs")
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   634
proof
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
  assume ?lhs
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   636
  then obtain S T
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   637
    where st: "open S" "z \<in> S" "f holomorphic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   638
              "open T" "z \<in> T" "g holomorphic_on T"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
    by (auto simp: analytic_at)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   640
  then show ?rhs
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   641
    by (metis Int_iff holomorphic_on_subset inf_le1 inf_le2 open_Int)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
next
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   643
  assume ?rhs
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   644
  then show ?lhs
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   645
    by (force simp add: analytic_at)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   648
subsection\<^marker>\<open>tag unimportant\<close>\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   650
lemma
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
  assumes "f analytic_on {z}" "g analytic_on {z}"
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   652
  shows complex_derivative_add_at: "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   653
    and complex_derivative_diff_at: "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   654
    and complex_derivative_mult_at: "deriv (\<lambda>w. f w * g w) z =
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   655
           f z * deriv g z + deriv f z * g z"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
proof -
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   657
  show "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   658
    using analytic_on_imp_differentiable_at assms by auto
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   659
  show "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   660
    using analytic_on_imp_differentiable_at assms by force
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   661
  obtain S where "open S" "z \<in> S" "f holomorphic_on S" "g holomorphic_on S"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   662
    using assms by (metis analytic_at_two)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   663
  then show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   664
    by (simp add: DERIV_imp_deriv [OF DERIV_mult'] holomorphic_derivI)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   667
lemma deriv_cmult_at:
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   668
  "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. c * f w) z = c * deriv f z"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   669
  by (auto simp: complex_derivative_mult_at)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   670
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   671
lemma deriv_cmult_right_at:
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   672
  "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. f w * c) z = deriv f z * c"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   673
  by (auto simp: complex_derivative_mult_at)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   675
subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex differentiation of sequences and series\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 61520
diff changeset
   677
(* TODO: Could probably be simplified using Uniform_Limit *)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
lemma has_complex_derivative_sequence:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   679
  fixes S :: "complex set"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   680
  assumes cvs: "convex S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   681
      and df:  "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x within S)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   682
      and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> S \<longrightarrow> norm (f' n x - g' x) \<le> e"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   683
      and "\<exists>x l. x \<in> S \<and> ((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   684
    shows "\<exists>g. \<forall>x \<in> S. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and>
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   685
                       (g has_field_derivative (g' x)) (at x within S)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   686
proof -
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   687
  from assms obtain x l where x: "x \<in> S" and tf: "((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   688
    by blast
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   689
  show ?thesis
68055
2cab37094fc4 more defer/prefer
paulson <lp15@cam.ac.uk>
parents: 67979
diff changeset
   690
    unfolding has_field_derivative_def
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
  proof (rule has_derivative_sequence [OF cvs _ _ x])
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68055
diff changeset
   692
    show "(\<lambda>n. f n x) \<longlonglongrightarrow> l"
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68055
diff changeset
   693
      by (rule tf)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   694
  next 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   695
    have **: "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> \<epsilon> * cmod h"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   696
      if "\<epsilon> > 0" for \<epsilon>::real 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   697
      by (metis that left_diff_distrib mult_right_mono norm_ge_zero norm_mult conv)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   698
    show "\<And>e. e > 0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68055
diff changeset
   699
      unfolding eventually_sequentially by (blast intro: **)
68055
2cab37094fc4 more defer/prefer
paulson <lp15@cam.ac.uk>
parents: 67979
diff changeset
   700
  qed (metis has_field_derivative_def df)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   701
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   703
lemma has_complex_derivative_series:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   704
  fixes S :: "complex set"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   705
  assumes cvs: "convex S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   706
      and df:  "\<And>n x. x \<in> S \<Longrightarrow> (f n has_field_derivative f' n x) (at x within S)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   707
      and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> S
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   708
                \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   709
      and "\<exists>x l. x \<in> S \<and> ((\<lambda>n. f n x) sums l)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   710
    shows "\<exists>g. \<forall>x \<in> S. ((\<lambda>n. f n x) sums g x) \<and> ((g has_field_derivative g' x) (at x within S))"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   711
proof -
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   712
  from assms obtain x l where x: "x \<in> S" and sf: "((\<lambda>n. f n x) sums l)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
    by blast
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   714
  { fix \<epsilon>::real assume e: "\<epsilon> > 0"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   715
    then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> S
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   716
            \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> \<epsilon>"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   717
      by (metis conv)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   718
    have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>S. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> \<epsilon> * cmod h"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   719
    proof (rule exI [of _ N], clarify)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   720
      fix n y h
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   721
      assume "N \<le> n" "y \<in> S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   722
      have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * \<epsilon>"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   723
        by (simp add: N \<open>N \<le> n\<close> \<open>y \<in> S\<close> mult_le_cancel_left)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   724
      then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> \<epsilon> * cmod h"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   725
        by (simp add: norm_mult [symmetric] field_simps sum_distrib_left)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   726
    qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   727
  } note ** = this
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   728
  show ?thesis
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   729
  unfolding has_field_derivative_def
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
  proof (rule has_derivative_series [OF cvs _ _ x])
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
    fix n x
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   732
    assume "x \<in> S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   733
    then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within S)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
      by (metis df has_field_derivative_def mult_commute_abs)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
  next show " ((\<lambda>n. f n x) sums l)"
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   736
    by (rule sf)
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   737
  next show "\<And>e. e>0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>S. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"
68239
0764ee22a4d1 tidy up of Derivative
paulson <lp15@cam.ac.uk>
parents: 68055
diff changeset
   738
      unfolding eventually_sequentially by (blast intro: **)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   739
  qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69529
diff changeset
   742
subsection\<^marker>\<open>tag unimportant\<close> \<open>Taylor on Complex Numbers\<close>
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   744
lemma sum_Suc_reindex:
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   745
  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   746
  shows "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   747
  by (induct n) auto
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   749
lemma field_Taylor:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   750
  assumes S: "convex S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   751
      and f: "\<And>i x. x \<in> S \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within S)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   752
      and B: "\<And>x. x \<in> S \<Longrightarrow> norm (f (Suc n) x) \<le> B"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   753
      and w: "w \<in> S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   754
      and z: "z \<in> S"
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   755
    shows "norm(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i)))
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   756
          \<le> B * norm(z - w)^(Suc n) / fact n"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
proof -
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   758
  have wzs: "closed_segment w z \<subseteq> S" using assms
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
    by (metis convex_contains_segment)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
  { fix u
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
    assume "u \<in> closed_segment w z"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   762
    then have "u \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
      by (metis wzs subsetD)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   764
    have *: "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) +
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   765
                      f (Suc i) u * (z-u)^i / (fact i)) =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   766
              f (Suc n) u * (z-u) ^ n / (fact n)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
    proof (induction n)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
      case 0 show ?case by simp
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
    next
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
      case (Suc n)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   771
      have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) +
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   772
                             f (Suc i) u * (z-u) ^ i / (fact i)) =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   773
           f (Suc n) u * (z-u) ^ n / (fact n) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   774
           f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) -
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   775
           f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56409
diff changeset
   776
        using Suc by simp
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   777
      also have "\<dots> = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
      proof -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   779
        have "(fact(Suc n)) *
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   780
             (f(Suc n) u *(z-u) ^ n / (fact n) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   781
               f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) -
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   782
               f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   783
            ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   784
            ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) -
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   785
            ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63332
diff changeset
   786
          by (simp add: algebra_simps del: fact_Suc)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   787
        also have "\<dots> = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) +
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   788
                         (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   789
                         (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63332
diff changeset
   790
          by (simp del: fact_Suc)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   791
        also have "\<dots> = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   792
                         (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   793
                         (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63332
diff changeset
   794
          by (simp only: fact_Suc of_nat_mult ac_simps) simp
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   795
        also have "\<dots> = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
          by (simp add: algebra_simps)
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   797
        finally show ?thesis
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63332
diff changeset
   798
        by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   799
      qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   800
      finally show ?case .
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   801
    qed
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   802
    have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i)))
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   803
                has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n))
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   804
               (at u within S)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   805
      unfolding * [symmetric]
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   806
      by (rule derivative_eq_intros assms \<open>u \<in> S\<close> refl | auto simp: field_simps)+
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   807
  } note sum_deriv = this
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
  { fix u
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
    assume u: "u \<in> closed_segment w z"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   810
    then have us: "u \<in> S"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
      by (metis wzs subsetD)
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   812
    have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> norm (f (Suc n) u) * norm (u - z) ^ n"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
      by (metis norm_minus_commute order_refl)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   814
    also have "\<dots> \<le> norm (f (Suc n) u) * norm (z - w) ^ n"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   815
      by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   816
    also have "\<dots> \<le> B * norm (z - w) ^ n"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
      by (metis norm_ge_zero zero_le_power mult_right_mono  B [OF us])
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   818
    finally have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> B * norm (z - w) ^ n" .
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
  } note cmod_bound = this
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   820
  have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)) = (\<Sum>i\<le>n. (f i z / (fact i)) * 0 ^ i)"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
    by simp
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   822
  also have "\<dots> = f 0 z / (fact 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   823
    by (subst sum_zero_power) simp
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   824
  finally have "norm (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)))
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   825
                \<le> norm ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   826
                        (\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))"
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   827
    by (simp add: norm_minus_commute)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   828
  also have "\<dots> \<le> B * norm (z - w) ^ n / (fact n) * norm (w - z)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   829
  proof (rule field_differentiable_bound)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   830
    show "\<And>x. x \<in> closed_segment w z \<Longrightarrow>
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   831
          ((\<lambda>\<xi>. \<Sum>i\<le>n. f i \<xi> * (z - \<xi>) ^ i / fact i) has_field_derivative f (Suc n) x * (z - x) ^ n / fact n)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   832
           (at x within closed_segment w z)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   833
      using DERIV_subset sum_deriv wzs by blast
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   834
  qed (auto simp: norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   835
  also have "\<dots>  \<le> B * norm (z - w) ^ Suc n / (fact n)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   836
    by (simp add: algebra_simps norm_minus_commute)
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  finally show ?thesis .
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
qed
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   839
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   840
lemma complex_Taylor:
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   841
  assumes S: "convex S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   842
      and f: "\<And>i x. x \<in> S \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within S)"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   843
      and B: "\<And>x. x \<in> S \<Longrightarrow> cmod (f (Suc n) x) \<le> B"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   844
      and w: "w \<in> S"
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 68239
diff changeset
   845
      and z: "z \<in> S"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   846
    shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) \<le> B * cmod(z - w)^(Suc n) / fact n"
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   847
  using assms by (rule field_Taylor)
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   848
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 66089
diff changeset
   849
62408
86f27b264d3d Conformal_mappings: a big development in complex analysis (+ some lemmas)
paulson <lp15@cam.ac.uk>
parents: 62397
diff changeset
   850
text\<open>Something more like the traditional MVT for real components\<close>
56370
7c717ba55a0b reorder Complex_Analysis_Basics; rename DD to deriv
hoelzl
parents: 56369
diff changeset
   851
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   852
lemma complex_mvt_line:
56369
2704ca85be98 moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
hoelzl
parents: 56332
diff changeset
   853
  assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61235
diff changeset
   854
    shows "\<exists>u. u \<in> closed_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))"
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   855
proof -
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   856
  define \<phi> where "\<phi> \<equiv> \<lambda>t. (1 - t) *\<^sub>R w + t *\<^sub>R z"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   857
  have twz: "\<And>t. \<phi> t = w + t *\<^sub>R (z - w)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   858
    by (simp add: \<phi>_def real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)
56381
0556204bc230 merged DERIV_intros, has_derivative_intros into derivative_intros
hoelzl
parents: 56371
diff changeset
   859
  note assms[unfolded has_field_derivative_def, derivative_intros]
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   860
  have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   861
        \<Longrightarrow> (Re \<circ> f \<circ> \<phi> has_derivative Re \<circ> (*) (f' (\<phi> x)) \<circ> (\<lambda>t. t *\<^sub>R (z - w)))
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   862
            (at x within {0..1})"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   863
    unfolding \<phi>_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   864
    by (intro derivative_eq_intros has_derivative_at_withinI) 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   865
       (auto simp: in_segment scaleR_right_diff_distrib)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   866
  obtain x where "0<x" "x<1" "(Re \<circ> f \<circ> \<phi>) 1 -
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   867
       (Re \<circ> f \<circ> \<phi>) 0 = (Re \<circ> (*) (f' (\<phi> x)) \<circ> (\<lambda>t. t *\<^sub>R (z - w))) (1 - 0)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   868
    using mvt_simple [OF zero_less_one *] by force
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   869
  then show ?thesis
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   870
    unfolding \<phi>_def
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   871
    by (smt (verit) comp_apply in_segment(1) scaleR_left_distrib scaleR_one scaleR_zero_left)
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   872
qed
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   873
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   874
lemma complex_Taylor_mvt:
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   875
  assumes "\<And>i x. \<lbrakk>x \<in> closed_segment w z; i \<le> n\<rbrakk> \<Longrightarrow> ((f i) has_field_derivative f (Suc i) x) (at x)"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   876
    shows "\<exists>u. u \<in> closed_segment w z \<and>
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   877
            Re (f 0 z) =
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   878
            Re ((\<Sum>i = 0..n. f i w * (z - w) ^ i / (fact i)) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   879
                (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))"
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   880
proof -
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   881
  { fix u
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   882
    assume u: "u \<in> closed_segment w z"
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   883
    have "(\<Sum>i = 0..n.
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   884
               (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) /
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   885
               (fact i)) =
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   886
          f (Suc 0) u -
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   887
             (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   888
             (fact (Suc n)) +
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   889
             (\<Sum>i = 0..n.
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   890
                 (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) /
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   891
                 (fact (Suc i)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   892
       by (subst sum_Suc_reindex) simp
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   893
    also have "\<dots> = f (Suc 0) u -
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   894
             (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   895
             (fact (Suc n)) +
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   896
             (\<Sum>i = 0..n.
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61531
diff changeset
   897
                 f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i))  -
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   898
                 f (Suc i) u * (z-u) ^ i / (fact i))"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 56889
diff changeset
   899
      by (simp only: diff_divide_distrib fact_cancel ac_simps)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   900
    also have "\<dots> = f (Suc 0) u -
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   901
             (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) /
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   902
             (fact (Suc n)) +
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   903
             f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63941
diff changeset
   904
      by (subst sum_Suc_diff) auto
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   905
    also have "\<dots> = f (Suc n) u * (z-u) ^ n / (fact n)"
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   906
      by (simp only: algebra_simps diff_divide_distrib fact_cancel)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   907
    finally have *: "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   908
                             - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) =
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   909
                  f (Suc n) u * (z - u) ^ n / (fact n)" .
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   910
    have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   911
                f (Suc n) u * (z - u) ^ n / (fact n))  (at u)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   912
      unfolding * [symmetric]
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   913
      by (rule derivative_eq_intros assms u refl | auto simp: field_simps)+
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   914
  }
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   915
  then show ?thesis
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   916
    apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / (fact i)"
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59615
diff changeset
   917
               "\<lambda>u. (f (Suc n) u * (z-u)^n / (fact n))"])
56238
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   918
    apply (auto simp add: intro: open_closed_segment)
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   919
    done
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   920
qed
5d147e1e18d1 a few new lemmas and generalisations of old ones
paulson <lp15@cam.ac.uk>
parents: 56223
diff changeset
   921
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59730
diff changeset
   922
56215
fcf90317383d New complex analysis material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
end