src/HOL/Analysis/Complex_Analysis_Basics.thy
changeset 68721 53ad5c01be3f
parent 68296 69d680e94961
child 69064 5840724b1d71
equal deleted inserted replaced
68720:1e1818612124 68721:53ad5c01be3f
    38   fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)"
    38   fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)"
    39   by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI)
    39   by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI)
    40 
    40 
    41 lemma linear_cnj: "linear cnj"
    41 lemma linear_cnj: "linear cnj"
    42   using bounded_linear.linear[OF bounded_linear_cnj] .
    42   using bounded_linear.linear[OF bounded_linear_cnj] .
       
    43 
       
    44 lemma vector_derivative_cnj_within:
       
    45   assumes "at x within A \<noteq> bot" and "f differentiable at x within A"
       
    46   shows   "vector_derivative (\<lambda>z. cnj (f z)) (at x within A) = 
       
    47              cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D")
       
    48 proof -
       
    49   let ?D = "vector_derivative f (at x within A)"
       
    50   from assms have "(f has_vector_derivative ?D) (at x within A)"
       
    51     by (subst (asm) vector_derivative_works)
       
    52   hence "((\<lambda>x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)"
       
    53     by (rule has_vector_derivative_cnj)
       
    54   thus ?thesis using assms by (auto dest: vector_derivative_within)
       
    55 qed
       
    56 
       
    57 lemma vector_derivative_cnj:
       
    58   assumes "f differentiable at x"
       
    59   shows   "vector_derivative (\<lambda>z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))"
       
    60   using assms by (intro vector_derivative_cnj_within) auto
    43 
    61 
    44 lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = ( * ) 0"
    62 lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = ( * ) 0"
    45   by auto
    63   by auto
    46 
    64 
    47 lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = ( * ) 1"
    65 lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = ( * ) 1"
   283   by (auto simp: holomorphic_on_def)
   301   by (auto simp: holomorphic_on_def)
   284 
   302 
   285 lemma holomorphic_on_compose_gen:
   303 lemma holomorphic_on_compose_gen:
   286   "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s"
   304   "f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s"
   287   by (metis holomorphic_on_compose holomorphic_on_subset)
   305   by (metis holomorphic_on_compose holomorphic_on_subset)
       
   306 
       
   307 lemma holomorphic_on_balls_imp_entire:
       
   308   assumes "\<not>bdd_above A" "\<And>r. r \<in> A \<Longrightarrow> f holomorphic_on ball c r"
       
   309   shows   "f holomorphic_on B"
       
   310 proof (rule holomorphic_on_subset)
       
   311   show "f holomorphic_on UNIV" unfolding holomorphic_on_def
       
   312   proof
       
   313     fix z :: complex
       
   314     from \<open>\<not>bdd_above A\<close> obtain r where r: "r \<in> A" "r > norm (z - c)"
       
   315       by (meson bdd_aboveI not_le)
       
   316     with assms(2) have "f holomorphic_on ball c r" by blast
       
   317     moreover from r have "z \<in> ball c r" by (auto simp: dist_norm norm_minus_commute)
       
   318     ultimately show "f field_differentiable at z"
       
   319       by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"])
       
   320   qed
       
   321 qed auto
       
   322 
       
   323 lemma holomorphic_on_balls_imp_entire':
       
   324   assumes "\<And>r. r > 0 \<Longrightarrow> f holomorphic_on ball c r"
       
   325   shows   "f holomorphic_on B"
       
   326 proof (rule holomorphic_on_balls_imp_entire)
       
   327   {
       
   328     fix M :: real
       
   329     have "\<exists>x. x > max M 0" by (intro gt_ex)
       
   330     hence "\<exists>x>0. x > M" by auto
       
   331   }
       
   332   thus "\<not>bdd_above {(0::real)<..}" unfolding bdd_above_def
       
   333     by (auto simp: not_le)
       
   334 qed (insert assms, auto)
   288 
   335 
   289 lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s"
   336 lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s"
   290   by (metis field_differentiable_minus holomorphic_on_def)
   337   by (metis field_differentiable_minus holomorphic_on_def)
   291 
   338 
   292 lemma holomorphic_on_add [holomorphic_intros]:
   339 lemma holomorphic_on_add [holomorphic_intros]: