src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy
author paulson <lp15@cam.ac.uk>
Mon, 07 Dec 2015 16:44:26 +0000
changeset 61806 d2e62ae01cd8
parent 61762 d50b993b4fb9
child 61808 fc1556774cfe
permissions -rw-r--r--
Cauchy's integral formula for circles. Starting to fix eventually_mono.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
section \<open>Complex path integrals and Cauchy's integral theorem\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
     3
text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
     4
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
theory Cauchy_Integral_Thm
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
     6
imports Complex_Transcendental Weierstrass Ordered_Euclidean_Space
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
begin
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
     9
subsection \<open>Piecewise differentiable functions\<close>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
definition piecewise_differentiable_on
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
           (infixr "piecewise'_differentiable'_on" 50)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
  where "f piecewise_differentiable_on i  \<equiv>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
           continuous_on i f \<and>
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    15
           (\<exists>s. finite s \<and> (\<forall>x \<in> i - s. f differentiable (at x within i)))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
lemma piecewise_differentiable_on_imp_continuous_on:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
    "f piecewise_differentiable_on s \<Longrightarrow> continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
by (simp add: piecewise_differentiable_on_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
lemma piecewise_differentiable_on_subset:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
    "f piecewise_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_differentiable_on t"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    23
  using continuous_on_subset
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    24
  unfolding piecewise_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    25
  apply safe
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    26
  apply (blast intro: elim: continuous_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    27
  by (meson Diff_iff differentiable_within_subset subsetCE)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    29
lemma differentiable_on_imp_piecewise_differentiable:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  fixes a:: "'a::{linorder_topology,real_normed_vector}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  shows "f differentiable_on {a..b} \<Longrightarrow> f piecewise_differentiable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    32
  apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    33
  apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    34
  done
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
lemma differentiable_imp_piecewise_differentiable:
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    37
    "(\<And>x. x \<in> s \<Longrightarrow> f differentiable (at x within s))
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
         \<Longrightarrow> f piecewise_differentiable_on s"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    39
by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    40
         intro: differentiable_within_subset)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
61204
3e491e34a62e new lemmas and movement of lemmas into place
paulson
parents: 61200
diff changeset
    42
lemma piecewise_differentiable_const [iff]: "(\<lambda>x. z) piecewise_differentiable_on s"
3e491e34a62e new lemmas and movement of lemmas into place
paulson
parents: 61200
diff changeset
    43
  by (simp add: differentiable_imp_piecewise_differentiable)
3e491e34a62e new lemmas and movement of lemmas into place
paulson
parents: 61200
diff changeset
    44
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
lemma piecewise_differentiable_compose:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
    "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s);
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
      \<Longrightarrow> (g o f) piecewise_differentiable_on s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
  apply (simp add: piecewise_differentiable_on_def, safe)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
  apply (blast intro: continuous_on_compose2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
  apply (rename_tac A B)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    53
  apply (blast intro: differentiable_chain_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    54
  done
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
lemma piecewise_differentiable_affine:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  fixes m::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  assumes "f piecewise_differentiable_on ((\<lambda>x. m *\<^sub>R x + c) ` s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_differentiable_on s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
proof (cases "m = 0")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
    unfolding o_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
    by (force intro: differentiable_imp_piecewise_differentiable differentiable_const)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
    apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
    apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
lemma piecewise_differentiable_cases:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  fixes c::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
  assumes "f piecewise_differentiable_on {a..c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
          "g piecewise_differentiable_on {c..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
           "a \<le> c" "c \<le> b" "f c = g c"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_differentiable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  obtain s t where st: "finite s" "finite t"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    81
                       "\<forall>x\<in>{a..c} - s. f differentiable at x within {a..c}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    82
                       "\<forall>x\<in>{c..b} - t. g differentiable at x within {c..b}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
    using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
    by (auto simp: piecewise_differentiable_on_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    85
  have finabc: "finite ({a,b,c} \<union> (s \<union> t))"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
    86
    by (metis \<open>finite s\<close> \<open>finite t\<close> finite_Un finite_insert finite.emptyI)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  have "continuous_on {a..c} f" "continuous_on {c..b} g"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    using assms piecewise_differentiable_on_def by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
  then have "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
                               OF closed_real_atLeastAtMost [of c b],
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
                               of f g "\<lambda>x. x\<le>c"]  assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
    by (force simp: ivl_disj_un_two_touch)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  moreover
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  { fix x
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    96
    assume x: "x \<in> {a..b} - ({a,b,c} \<union> (s \<union> t))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
    97
    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg")
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
    proof (cases x c rule: le_cases)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
      case le show ?diff_fg
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   100
        apply (rule differentiable_transform_within [where d = "dist x c" and f = f])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   101
        using x le st
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   102
        apply (simp_all add: dist_real_def dist_nz [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   103
        apply (rule differentiable_at_withinI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   104
        apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   105
        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   106
        apply (force elim!: differentiable_subset)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
      case ge show ?diff_fg
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   110
        apply (rule differentiable_transform_within [where d = "dist x c" and f = g])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   111
        using x ge st
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   112
        apply (simp_all add: dist_real_def dist_nz [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   113
        apply (rule differentiable_at_withinI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   114
        apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   115
        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   116
        apply (force elim!: differentiable_subset)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
    qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  }
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   120
  then have "\<exists>s. finite s \<and>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   121
                 (\<forall>x\<in>{a..b} - s. (\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b})"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   122
    by (meson finabc)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
  ultimately show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
    by (simp add: piecewise_differentiable_on_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
lemma piecewise_differentiable_neg:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    "f piecewise_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_differentiable_on s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  by (auto simp: piecewise_differentiable_on_def continuous_on_minus)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
lemma piecewise_differentiable_add:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
  assumes "f piecewise_differentiable_on i"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
          "g piecewise_differentiable_on i"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
    shows "(\<lambda>x. f x + g x) piecewise_differentiable_on i"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  obtain s t where st: "finite s" "finite t"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   137
                       "\<forall>x\<in>i - s. f differentiable at x within i"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   138
                       "\<forall>x\<in>i - t. g differentiable at x within i"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
    using assms by (auto simp: piecewise_differentiable_on_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   140
  then have "finite (s \<union> t) \<and> (\<forall>x\<in>i - (s \<union> t). (\<lambda>x. f x + g x) differentiable at x within i)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
    by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  moreover have "continuous_on i f" "continuous_on i g"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    using assms piecewise_differentiable_on_def by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  ultimately show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
    by (auto simp: piecewise_differentiable_on_def continuous_on_add)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
lemma piecewise_differentiable_diff:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
    "\<lbrakk>f piecewise_differentiable_on s;  g piecewise_differentiable_on s\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_differentiable_on s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
  unfolding diff_conv_add_uminus
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  by (metis piecewise_differentiable_add piecewise_differentiable_neg)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   154
lemma continuous_on_joinpaths_D1:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   155
    "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   156
  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   157
  apply (rule continuous_intros | simp)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   158
  apply (auto elim!: continuous_on_subset simp: joinpaths_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   159
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   160
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   161
lemma continuous_on_joinpaths_D2:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   162
    "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   163
  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\<lambda>x. inverse 2*x + 1/2)"])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   164
  apply (rule continuous_intros | simp)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   165
  apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   166
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   167
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   168
lemma piecewise_differentiable_D1:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   169
    "(g1 +++ g2) piecewise_differentiable_on {0..1} \<Longrightarrow> g1 piecewise_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   170
  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   171
  apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   172
  apply simp
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   173
  apply (intro ballI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   174
  apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   175
       in differentiable_transform_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   176
  apply (auto simp: dist_real_def joinpaths_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   177
  apply (rule differentiable_chain_within derivative_intros | simp)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   178
  apply (rule differentiable_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   179
  apply (force simp:)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   180
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   181
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   182
lemma piecewise_differentiable_D2:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   183
    "\<lbrakk>(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\<rbrakk>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   184
    \<Longrightarrow> g2 piecewise_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   185
  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   186
  apply (rule_tac x="insert 0 ((\<lambda>x. 2*x-1)`s)" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   187
  apply simp
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   188
  apply (intro ballI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   189
  apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   190
          in differentiable_transform_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   191
  apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: split_if_asm)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   192
  apply (rule differentiable_chain_within derivative_intros | simp)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   193
  apply (rule differentiable_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   194
  apply (force simp: divide_simps)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   195
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   196
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   197
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   198
subsubsection\<open>The concept of continuously differentiable\<close>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   199
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   200
definition C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   201
           (infix "C1'_differentiable'_on" 50)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   202
  where
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   203
  "f C1_differentiable_on s \<longleftrightarrow>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   204
   (\<exists>D. (\<forall>x \<in> s. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on s D)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   205
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   206
lemma C1_differentiable_on_eq:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   207
    "f C1_differentiable_on s \<longleftrightarrow>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   208
     (\<forall>x \<in> s. f differentiable at x) \<and> continuous_on s (\<lambda>x. vector_derivative f (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   209
  unfolding C1_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   210
  apply safe
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   211
  using differentiable_def has_vector_derivative_def apply blast
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   212
  apply (erule continuous_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   213
  using vector_derivative_at apply fastforce
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   214
  using vector_derivative_works apply fastforce
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   215
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   216
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   217
lemma C1_differentiable_on_subset:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   218
  "f C1_differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   219
  unfolding C1_differentiable_on_def  continuous_on_eq_continuous_within
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   220
  by (blast intro:  continuous_within_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   221
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   222
lemma C1_differentiable_compose:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   223
    "\<lbrakk>f C1_differentiable_on s; g C1_differentiable_on (f ` s);
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   224
      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   225
      \<Longrightarrow> (g o f) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   226
  apply (simp add: C1_differentiable_on_eq, safe)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   227
   using differentiable_chain_at apply blast
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   228
  apply (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   229
   apply (rule Limits.continuous_on_scaleR, assumption)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   230
   apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   231
  by (simp add: vector_derivative_chain_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   232
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   233
lemma C1_diff_imp_diff: "f C1_differentiable_on s \<Longrightarrow> f differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   234
  by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   235
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   236
lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\<lambda>x. x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   237
  by (auto simp: C1_differentiable_on_eq continuous_on_const)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   238
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   239
lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   240
  by (auto simp: C1_differentiable_on_eq continuous_on_const)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   241
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   242
lemma C1_differentiable_on_add [simp, derivative_intros]:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   243
  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x + g x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   244
  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   245
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   246
lemma C1_differentiable_on_minus [simp, derivative_intros]:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   247
  "f C1_differentiable_on s \<Longrightarrow> (\<lambda>x. - f x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   248
  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   249
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   250
lemma C1_differentiable_on_diff [simp, derivative_intros]:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   251
  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x - g x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   252
  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   253
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   254
lemma C1_differentiable_on_mult [simp, derivative_intros]:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   255
  fixes f g :: "real \<Rightarrow> 'a :: real_normed_algebra"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   256
  shows "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x * g x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   257
  unfolding C1_differentiable_on_eq
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   258
  by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   259
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   260
lemma C1_differentiable_on_scaleR [simp, derivative_intros]:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   261
  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   262
  unfolding C1_differentiable_on_eq
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   263
  by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   264
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   265
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   266
definition piecewise_C1_differentiable_on
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   267
           (infixr "piecewise'_C1'_differentiable'_on" 50)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   268
  where "f piecewise_C1_differentiable_on i  \<equiv>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   269
           continuous_on i f \<and>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   270
           (\<exists>s. finite s \<and> (f C1_differentiable_on (i - s)))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   271
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   272
lemma C1_differentiable_imp_piecewise:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   273
    "f C1_differentiable_on s \<Longrightarrow> f piecewise_C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   274
  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   275
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   276
lemma piecewise_C1_imp_differentiable:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   277
    "f piecewise_C1_differentiable_on i \<Longrightarrow> f piecewise_differentiable_on i"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   278
  by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   279
           C1_differentiable_on_def differentiable_def has_vector_derivative_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   280
           intro: has_derivative_at_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   281
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   282
lemma piecewise_C1_differentiable_compose:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   283
    "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s);
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   284
      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   285
      \<Longrightarrow> (g o f) piecewise_C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   286
  apply (simp add: piecewise_C1_differentiable_on_def, safe)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   287
  apply (blast intro: continuous_on_compose2)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   288
  apply (rename_tac A B)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   289
  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   290
  apply (rule conjI, blast)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   291
  apply (rule C1_differentiable_compose)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   292
  apply (blast intro: C1_differentiable_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   293
  apply (blast intro: C1_differentiable_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   294
  by (simp add: Diff_Int_distrib2)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   295
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   296
lemma piecewise_C1_differentiable_on_subset:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   297
    "f piecewise_C1_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_C1_differentiable_on t"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   298
  by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   299
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   300
lemma C1_differentiable_imp_continuous_on:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   301
  "f C1_differentiable_on s \<Longrightarrow> continuous_on s f"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   302
  unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   303
  using differentiable_at_withinI differentiable_imp_continuous_within by blast
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   304
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   305
lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   306
  unfolding C1_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   307
  by auto
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   308
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   309
lemma piecewise_C1_differentiable_affine:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   310
  fixes m::real
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   311
  assumes "f piecewise_C1_differentiable_on ((\<lambda>x. m * x + c) ` s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   312
  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   313
proof (cases "m = 0")
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   314
  case True
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   315
  then show ?thesis
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   316
    unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   317
next
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   318
  case False
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   319
  show ?thesis
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   320
    apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   321
    apply (rule assms derivative_intros | simp add: False vimage_def)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   322
    using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real]
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   323
    apply simp
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   324
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   325
qed
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   326
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   327
lemma piecewise_C1_differentiable_cases:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   328
  fixes c::real
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   329
  assumes "f piecewise_C1_differentiable_on {a..c}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   330
          "g piecewise_C1_differentiable_on {c..b}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   331
           "a \<le> c" "c \<le> b" "f c = g c"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   332
  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_C1_differentiable_on {a..b}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   333
proof -
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   334
  obtain s t where st: "f C1_differentiable_on ({a..c} - s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   335
                       "g C1_differentiable_on ({c..b} - t)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   336
                       "finite s" "finite t"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   337
    using assms
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   338
    by (force simp: piecewise_C1_differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   339
  then have f_diff: "f differentiable_on {a..<c} - s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   340
        and g_diff: "g differentiable_on {c<..b} - t"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   341
    by (simp_all add: C1_differentiable_on_eq differentiable_at_withinI differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   342
  have "continuous_on {a..c} f" "continuous_on {c..b} g"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   343
    using assms piecewise_C1_differentiable_on_def by auto
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   344
  then have cab: "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   345
    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   346
                               OF closed_real_atLeastAtMost [of c b],
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   347
                               of f g "\<lambda>x. x\<le>c"]  assms
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   348
    by (force simp: ivl_disj_un_two_touch)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   349
  { fix x
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   350
    assume x: "x \<in> {a..b} - insert c (s \<union> t)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   351
    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x" (is "?diff_fg")
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   352
    proof (cases x c rule: le_cases)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   353
      case le show ?diff_fg
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   354
        apply (rule differentiable_transform_at [of "dist x c" _ f])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   355
        using x dist_real_def le st by (auto simp: C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   356
    next
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   357
      case ge show ?diff_fg
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   358
        apply (rule differentiable_transform_at [of "dist x c" _ g])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   359
        using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   360
    qed
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   361
  }
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   362
  then have "(\<forall>x \<in> {a..b} - insert c (s \<union> t). (\<lambda>x. if x \<le> c then f x else g x) differentiable at x)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   363
    by auto
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   364
  moreover
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   365
  { assume fcon: "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative f (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   366
       and gcon: "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative g (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   367
    have "open ({a<..<c} - s)"  "open ({c<..<b} - t)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   368
      using st by (simp_all add: open_Diff finite_imp_closed)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   369
    moreover have "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   370
      apply (rule continuous_on_eq [OF fcon])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   371
      apply (simp add:)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   372
      apply (rule vector_derivative_at [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   373
      apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   374
      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   375
      using f_diff
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   376
      by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_imp_le st(1))
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   377
    moreover have "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   378
      apply (rule continuous_on_eq [OF gcon])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   379
      apply (simp add:)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   380
      apply (rule vector_derivative_at [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   381
      apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   382
      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   383
      using g_diff
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   384
      by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_imp_le st(2))
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   385
    ultimately have "continuous_on ({a<..<b} - insert c (s \<union> t))
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   386
        (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   387
      apply (rule continuous_on_subset [OF continuous_on_open_Un], auto)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   388
      done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   389
  } note * = this
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   390
  have "continuous_on ({a<..<b} - insert c (s \<union> t)) (\<lambda>x. vector_derivative (\<lambda>x. if x \<le> c then f x else g x) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   391
    using st
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   392
    by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   393
  ultimately have "\<exists>s. finite s \<and> ((\<lambda>x. if x \<le> c then f x else g x) C1_differentiable_on {a..b} - s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   394
    apply (rule_tac x="{a,b,c} \<union> s \<union> t" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   395
    using st  by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   396
  with cab show ?thesis
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   397
    by (simp add: piecewise_C1_differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   398
qed
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   399
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   400
lemma piecewise_C1_differentiable_neg:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   401
    "f piecewise_C1_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   402
  unfolding piecewise_C1_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   403
  by (auto intro!: continuous_on_minus C1_differentiable_on_minus)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   404
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   405
lemma piecewise_C1_differentiable_add:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   406
  assumes "f piecewise_C1_differentiable_on i"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   407
          "g piecewise_C1_differentiable_on i"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   408
    shows "(\<lambda>x. f x + g x) piecewise_C1_differentiable_on i"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   409
proof -
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   410
  obtain s t where st: "finite s" "finite t"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   411
                       "f C1_differentiable_on (i-s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   412
                       "g C1_differentiable_on (i-t)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   413
    using assms by (auto simp: piecewise_C1_differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   414
  then have "finite (s \<union> t) \<and> (\<lambda>x. f x + g x) C1_differentiable_on i - (s \<union> t)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   415
    by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   416
  moreover have "continuous_on i f" "continuous_on i g"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   417
    using assms piecewise_C1_differentiable_on_def by auto
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   418
  ultimately show ?thesis
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   419
    by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   420
qed
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   421
61204
3e491e34a62e new lemmas and movement of lemmas into place
paulson
parents: 61200
diff changeset
   422
lemma piecewise_C1_differentiable_diff:
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   423
    "\<lbrakk>f piecewise_C1_differentiable_on s;  g piecewise_C1_differentiable_on s\<rbrakk>
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   424
     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   425
  unfolding diff_conv_add_uminus
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   426
  by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   427
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   428
lemma piecewise_C1_differentiable_D1:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   429
  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   430
  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   431
    shows "g1 piecewise_C1_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   432
proof -
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   433
  obtain s where "finite s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   434
             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   435
             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   436
    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   437
  then have g1D: "g1 differentiable at x" if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   438
    apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   439
    using that
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   440
    apply (simp_all add: dist_real_def joinpaths_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   441
    apply (rule differentiable_chain_at derivative_intros | force)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   442
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   443
  have [simp]: "vector_derivative (g1 \<circ> op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   444
               if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   445
    apply (subst vector_derivative_chain_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   446
    using that
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   447
    apply (rule derivative_eq_intros g1D | simp)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   448
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   449
  have "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   450
    using co12 by (rule continuous_on_subset) force
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   451
  then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 o op*2) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   452
    apply (rule continuous_on_eq [OF _ vector_derivative_at])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   453
    apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   454
    apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   455
    apply (force intro: g1D differentiable_chain_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   456
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   457
  have "continuous_on ({0..1} - insert 1 (op * 2 ` s))
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   458
                      ((\<lambda>x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   459
    apply (rule continuous_intros)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   460
    using coDhalf
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   461
    apply (simp add: scaleR_conv_of_real image_set_diff image_image)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   462
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   463
  then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\<lambda>x. vector_derivative g1 (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   464
    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   465
  have "continuous_on {0..1} g1"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   466
    using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
   467
  with \<open>finite s\<close> show ?thesis
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   468
    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   469
    apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   470
    apply (simp add: g1D con_g1)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   471
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   472
qed
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   473
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   474
lemma piecewise_C1_differentiable_D2:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   475
  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   476
  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   477
    shows "g2 piecewise_C1_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   478
proof -
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   479
  obtain s where "finite s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   480
             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   481
             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   482
    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   483
  then have g2D: "g2 differentiable at x" if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   484
    apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" in differentiable_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   485
    using that
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   486
    apply (simp_all add: dist_real_def joinpaths_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   487
    apply (auto simp: dist_real_def joinpaths_def field_simps)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   488
    apply (rule differentiable_chain_at derivative_intros | force)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   489
    apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   490
    apply assumption
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   491
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   492
  have [simp]: "vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   493
               if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   494
    using that  by (auto simp: vector_derivative_chain_at divide_simps g2D)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   495
  have "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   496
    using co12 by (rule continuous_on_subset) force
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   497
  then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g2 o (\<lambda>x. 2*x-1)) (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   498
    apply (rule continuous_on_eq [OF _ vector_derivative_at])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   499
    apply (rule_tac f="g2 o (\<lambda>x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   500
    apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric]
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   501
                intro!: g2D differentiable_chain_at)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   502
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   503
  have [simp]: "((\<lambda>x. (x + 1) / 2) ` ({0..1} - insert 0 ((\<lambda>x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   504
    apply (simp add: image_set_diff inj_on_def image_image)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   505
    apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   506
    done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   507
  have "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s))
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   508
                      ((\<lambda>x. 1/2 * vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at x)) o (\<lambda>x. (x+1)/2))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   509
    by (rule continuous_intros | simp add:  coDhalf)+
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   510
  then have con_g2: "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) (\<lambda>x. vector_derivative g2 (at x))"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   511
    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   512
  have "continuous_on {0..1} g2"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   513
    using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
   514
  with \<open>finite s\<close> show ?thesis
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   515
    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   516
    apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` s)" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   517
    apply (simp add: g2D con_g2)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   518
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   519
qed
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
subsection \<open>Valid paths, and their start and finish\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
lemma Diff_Un_eq: "A - (B \<union> C) = A - B - C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
  by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
definition valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   527
  where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
  where "closed_path g \<equiv> g 0 = g 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
subsubsection\<open>In particular, all results for paths apply\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
lemma valid_path_imp_path: "valid_path g \<Longrightarrow> path g"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   535
by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
lemma connected_valid_path_image: "valid_path g \<Longrightarrow> connected(path_image g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
  by (metis connected_path_image valid_path_imp_path)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
lemma compact_valid_path_image: "valid_path g \<Longrightarrow> compact(path_image g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
  by (metis compact_path_image valid_path_imp_path)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
lemma bounded_valid_path_image: "valid_path g \<Longrightarrow> bounded(path_image g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
  by (metis bounded_path_image valid_path_imp_path)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
lemma closed_valid_path_image: "valid_path g \<Longrightarrow> closed(path_image g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  by (metis closed_path_image valid_path_imp_path)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
subsection\<open>Contour Integrals along a path\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   552
text\<open>This definition is for complex numbers only, and does not generalise to line integrals in a vector field\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   553
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   554
text\<open>piecewise differentiable function on [0,1]\<close>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   555
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   556
definition has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   557
           (infixr "has'_contour'_integral" 50)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   558
  where "(f has_contour_integral i) g \<equiv>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   559
           ((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   560
            has_integral i) {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   561
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   562
definition contour_integrable_on
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   563
           (infixr "contour'_integrable'_on" 50)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   564
  where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   565
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   566
definition contour_integral
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   567
  where "contour_integral g f \<equiv> @i. (f has_contour_integral i) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   568
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   569
lemma contour_integral_unique: "(f has_contour_integral i)  g \<Longrightarrow> contour_integral g f = i"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   570
  by (auto simp: contour_integral_def has_contour_integral_def integral_def [symmetric])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   571
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   572
lemma has_contour_integral_integral:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   573
    "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   574
  by (metis contour_integral_unique contour_integrable_on_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   575
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   576
lemma has_contour_integral_unique:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   577
    "(f has_contour_integral i) g \<Longrightarrow> (f has_contour_integral j) g \<Longrightarrow> i = j"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   578
  using has_integral_unique
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   579
  by (auto simp: has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   580
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   581
lemma has_contour_integral_integrable: "(f has_contour_integral i) g \<Longrightarrow> f contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   582
  using contour_integrable_on_def by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   583
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   584
(* Show that we can forget about the localized derivative.*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   585
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   586
lemma vector_derivative_within_interior:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   587
     "\<lbrakk>x \<in> interior s; NO_MATCH UNIV s\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   588
      \<Longrightarrow> vector_derivative f (at x within s) = vector_derivative f (at x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   589
  apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   590
  apply (subst lim_within_interior, auto)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   591
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   592
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   593
lemma has_integral_localized_vector_derivative:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   594
    "((\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \<longleftrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   595
     ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   596
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   597
  have "{a..b} - {a,b} = interior {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   598
    by (simp add: atLeastAtMost_diff_ends)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   599
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   600
    apply (rule has_integral_spike_eq [of "{a,b}"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   601
    apply (auto simp: vector_derivative_within_interior)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   602
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   603
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   604
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   605
lemma integrable_on_localized_vector_derivative:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   606
    "(\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \<longleftrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   607
     (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   608
  by (simp add: integrable_on_def has_integral_localized_vector_derivative)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   609
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   610
lemma has_contour_integral:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   611
     "(f has_contour_integral i) g \<longleftrightarrow>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   612
      ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   613
  by (simp add: has_integral_localized_vector_derivative has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   614
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   615
lemma contour_integrable_on:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   616
     "f contour_integrable_on g \<longleftrightarrow>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   617
      (\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   618
  by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   619
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   620
subsection\<open>Reversing a path\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   621
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   622
lemma valid_path_imp_reverse:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   623
  assumes "valid_path g"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   624
    shows "valid_path(reversepath g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   625
proof -
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   626
  obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   627
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   628
  then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   629
    apply (auto simp: reversepath_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   630
    apply (rule C1_differentiable_compose [of "\<lambda>x::real. 1-x" _ g, unfolded o_def])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   631
    apply (auto simp: C1_differentiable_on_eq)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   632
    apply (rule continuous_intros, force)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   633
    apply (force elim!: continuous_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   634
    apply (simp add: finite_vimageI inj_on_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   635
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   636
  then show ?thesis using assms
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   637
    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   638
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   639
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   640
lemma valid_path_reversepath: "valid_path(reversepath g) \<longleftrightarrow> valid_path g"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   641
  using valid_path_imp_reverse by force
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   642
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   643
lemma has_contour_integral_reversepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   644
  assumes "valid_path g" "(f has_contour_integral i) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   645
    shows "(f has_contour_integral (-i)) (reversepath g)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   646
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   647
  { fix s x
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   648
    assume xs: "g C1_differentiable_on ({0..1} - s)" "x \<notin> op - 1 ` s" "0 \<le> x" "x \<le> 1"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   649
      have "vector_derivative (\<lambda>x. g (1 - x)) (at x within {0..1}) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   650
            - vector_derivative g (at (1 - x) within {0..1})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   651
      proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   652
        obtain f' where f': "(g has_vector_derivative f') (at (1 - x))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   653
          using xs
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   654
          by (force simp: has_vector_derivative_def C1_differentiable_on_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   655
        have "(g o (\<lambda>x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   656
          apply (rule vector_diff_chain_within)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   657
          apply (intro vector_diff_chain_within derivative_eq_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   658
          apply (rule has_vector_derivative_at_within [OF f'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   659
          done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   660
        then have mf': "((\<lambda>x. g (1 - x)) has_vector_derivative -f') (at x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   661
          by (simp add: o_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   662
        show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   663
          using xs
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   664
          by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   665
      qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   666
  } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   667
  have 01: "{0..1::real} = cbox 0 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   668
    by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   669
  show ?thesis using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   670
    apply (auto simp: has_contour_integral_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   671
    apply (drule has_integral_affinity01 [where m= "-1" and c=1])
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   672
    apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   673
    apply (drule has_integral_neg)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   674
    apply (rule_tac s = "(\<lambda>x. 1 - x) ` s" in has_integral_spike_finite)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   675
    apply (auto simp: *)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   676
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   677
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   678
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   679
lemma contour_integrable_reversepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   680
    "valid_path g \<Longrightarrow> f contour_integrable_on g \<Longrightarrow> f contour_integrable_on (reversepath g)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   681
  using has_contour_integral_reversepath contour_integrable_on_def by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   682
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   683
lemma contour_integrable_reversepath_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   684
    "valid_path g \<Longrightarrow> (f contour_integrable_on (reversepath g) \<longleftrightarrow> f contour_integrable_on g)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   685
  using contour_integrable_reversepath valid_path_reversepath by fastforce
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   686
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   687
lemma contour_integral_reversepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   688
    "\<lbrakk>valid_path g; f contour_integrable_on g\<rbrakk> \<Longrightarrow> contour_integral (reversepath g) f = -(contour_integral g f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   689
  using has_contour_integral_reversepath contour_integrable_on_def contour_integral_unique by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   690
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   691
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   692
subsection\<open>Joining two paths together\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   693
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   694
lemma valid_path_join:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   695
  assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   696
    shows "valid_path(g1 +++ g2)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   697
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   698
  have "g1 1 = g2 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   699
    using assms by (auto simp: pathfinish_def pathstart_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   700
  moreover have "(g1 o (\<lambda>x. 2*x)) piecewise_C1_differentiable_on {0..1/2}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   701
    apply (rule piecewise_C1_differentiable_compose)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   702
    using assms
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   703
    apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   704
    apply (rule continuous_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   705
    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   706
    done
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   707
  moreover have "(g2 o (\<lambda>x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   708
    apply (rule piecewise_C1_differentiable_compose)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   709
    using assms unfolding valid_path_def piecewise_C1_differentiable_on_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   710
    by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   711
             simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   712
  ultimately show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   713
    apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   714
    apply (rule piecewise_C1_differentiable_cases)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   715
    apply (auto simp: o_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   716
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   717
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   718
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   719
lemma valid_path_join_D1:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   720
  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   721
  shows "valid_path (g1 +++ g2) \<Longrightarrow> valid_path g1"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   722
  unfolding valid_path_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   723
  by (rule piecewise_C1_differentiable_D1)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   724
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   725
lemma valid_path_join_D2:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   726
  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   727
  shows "\<lbrakk>valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> valid_path g2"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   728
  unfolding valid_path_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   729
  by (rule piecewise_C1_differentiable_D2)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   730
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   731
lemma valid_path_join_eq [simp]:
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   732
  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   733
  shows "pathfinish g1 = pathstart g2 \<Longrightarrow> (valid_path(g1 +++ g2) \<longleftrightarrow> valid_path g1 \<and> valid_path g2)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   734
  using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   735
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   736
lemma has_contour_integral_join:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   737
  assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   738
          "valid_path g1" "valid_path g2"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   739
    shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   740
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   741
  obtain s1 s2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   742
    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   743
      and s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   744
    using assms
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   745
    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   746
  have 1: "((\<lambda>x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   747
   and 2: "((\<lambda>x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   748
    using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   749
    by (auto simp: has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   750
  have i1: "((\<lambda>x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   751
   and i2: "((\<lambda>x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   752
    using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   753
          has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   754
    by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   755
  have g1: "\<lbrakk>0 \<le> z; z*2 < 1; z*2 \<notin> s1\<rbrakk> \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   756
            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   757
            2 *\<^sub>R vector_derivative g1 (at (z*2))" for z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   758
    apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "\<bar>z - 1/2\<bar>" _ "(\<lambda>x. g1(2*x))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   759
    apply (simp_all add: dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   760
    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x" 2 _ g1, simplified o_def])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   761
    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   762
    using s1
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   763
    apply (auto simp: algebra_simps vector_derivative_works)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   764
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   765
  have g2: "\<lbrakk>1 < z*2; z \<le> 1; z*2 - 1 \<notin> s2\<rbrakk> \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   766
            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   767
            2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   768
    apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "\<bar>z - 1/2\<bar>" _ "(\<lambda>x. g2 (2*x - 1))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   769
    apply (simp_all add: dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   770
    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x - 1" 2 _ g2, simplified o_def])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   771
    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   772
    using s2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   773
    apply (auto simp: algebra_simps vector_derivative_works)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   774
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   775
  have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   776
    apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   777
    using s1
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   778
    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   779
    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   780
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   781
  moreover have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   782
    apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\<lambda>x. 2*x-1) -` s2)"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   783
    using s2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   784
    apply (force intro: finite_vimageI [where h = "\<lambda>x. 2*x-1"] inj_onI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   785
    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   786
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   787
  ultimately
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   788
  show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   789
    apply (simp add: has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   790
    apply (rule has_integral_combine [where c = "1/2"], auto)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   791
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   792
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   793
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   794
lemma contour_integrable_joinI:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   795
  assumes "f contour_integrable_on g1" "f contour_integrable_on g2"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   796
          "valid_path g1" "valid_path g2"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   797
    shows "f contour_integrable_on (g1 +++ g2)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   798
  using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   799
  by (meson has_contour_integral_join contour_integrable_on_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   800
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   801
lemma contour_integrable_joinD1:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   802
  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   803
    shows "f contour_integrable_on g1"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   804
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   805
  obtain s1
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   806
    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   807
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   808
  have "(\<lambda>x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   809
    using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   810
    apply (auto simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   811
    apply (drule integrable_on_subcbox [where a=0 and b="1/2"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   812
    apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   813
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   814
  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   815
    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   816
  have g1: "\<lbrakk>0 < z; z < 1; z \<notin> s1\<rbrakk> \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   817
            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   818
            2 *\<^sub>R vector_derivative g1 (at z)"  for z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   819
    apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "\<bar>(z-1)/2\<bar>" _ "(\<lambda>x. g1(2*x))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   820
    apply (simp_all add: field_simps dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   821
    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2" 2 _ g1, simplified o_def])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   822
    using s1
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   823
    apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   824
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   825
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   826
    using s1
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   827
    apply (auto simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   828
    apply (rule integrable_spike_finite [of "{0,1} \<union> s1", OF _ _ *])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   829
    apply (auto simp: joinpaths_def scaleR_conv_of_real g1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   830
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   831
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   832
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   833
lemma contour_integrable_joinD2:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   834
  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   835
    shows "f contour_integrable_on g2"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   836
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   837
  obtain s2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   838
    where s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   839
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   840
  have "(\<lambda>x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   841
    using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   842
    apply (auto simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   843
    apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   844
    apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   845
    apply (simp add: image_affinity_atLeastAtMost_diff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   846
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   847
  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2)))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   848
                integrable_on {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   849
    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   850
  have g2: "\<lbrakk>0 < z; z < 1; z \<notin> s2\<rbrakk> \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   851
            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   852
            2 *\<^sub>R vector_derivative g2 (at z)" for z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   853
    apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "\<bar>z/2\<bar>" _ "(\<lambda>x. g2(2*x-1))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   854
    apply (simp_all add: field_simps dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   855
    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2-1" 2 _ g2, simplified o_def])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   856
    using s2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   857
    apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   858
                      vector_derivative_works add_divide_distrib)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   859
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   860
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   861
    using s2
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   862
    apply (auto simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   863
    apply (rule integrable_spike_finite [of "{0,1} \<union> s2", OF _ _ *])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   864
    apply (auto simp: joinpaths_def scaleR_conv_of_real g2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   865
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   866
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   867
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   868
lemma contour_integrable_join [simp]:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   869
  shows
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   870
    "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   871
     \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   872
using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   873
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   874
lemma contour_integral_join [simp]:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   875
  shows
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   876
    "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   877
        \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   878
  by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   879
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   880
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   881
subsection\<open>Shifting the starting point of a (closed) path\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   882
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   883
lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   884
  by (auto simp: shiftpath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   885
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   886
lemma valid_path_shiftpath [intro]:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   887
  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   888
    shows "valid_path(shiftpath a g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   889
  using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   890
  apply (auto simp: valid_path_def shiftpath_alt_def)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   891
  apply (rule piecewise_C1_differentiable_cases)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   892
  apply (auto simp: algebra_simps)
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   893
  apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   894
  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   895
  apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   896
  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   897
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   898
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   899
lemma has_contour_integral_shiftpath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   900
  assumes f: "(f has_contour_integral i) g" "valid_path g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   901
      and a: "a \<in> {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   902
    shows "(f has_contour_integral i) (shiftpath a g)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   903
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   904
  obtain s
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   905
    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   906
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   907
  have *: "((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   908
    using assms by (auto simp: has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   909
  then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   910
                    integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   911
    apply (rule has_integral_unique)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   912
    apply (subst add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   913
    apply (subst Integration.integral_combine)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   914
    using assms * integral_unique by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   915
  { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   916
    have "0 \<le> x \<Longrightarrow> x + a < 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a) ` s \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   917
         vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   918
      unfolding shiftpath_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   919
      apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "dist(1-a) x" _ "(\<lambda>x. g(a+x))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   920
        apply (auto simp: field_simps dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   921
      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a" 1 _ g, simplified o_def scaleR_one])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   922
       apply (intro derivative_eq_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   923
      using g
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   924
       apply (drule_tac x="x+a" in bspec)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   925
      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   926
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   927
  } note vd1 = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   928
  { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   929
    have "1 < x + a \<Longrightarrow> x \<le> 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a + 1) ` s \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   930
          vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   931
      unfolding shiftpath_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   932
      apply (rule vector_derivative_at [OF has_vector_derivative_transform_at [of "dist (1-a) x" _ "(\<lambda>x. g(a+x-1))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   933
        apply (auto simp: field_simps dist_real_def abs_if split: split_if_asm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   934
      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a-1" 1 _ g, simplified o_def scaleR_one])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   935
       apply (intro derivative_eq_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   936
      using g
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   937
      apply (drule_tac x="x+a-1" in bspec)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   938
      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   939
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   940
  } note vd2 = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   941
  have va1: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   942
    using * a   by (fastforce intro: integrable_subinterval_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   943
  have v0a: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   944
    apply (rule integrable_subinterval_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   945
    using * a by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   946
  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   947
        has_integral  integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {0..1 - a}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   948
    apply (rule has_integral_spike_finite
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   949
             [where s = "{1-a} \<union> (\<lambda>x. x-a) ` s" and f = "\<lambda>x. f(g(a+x)) * vector_derivative g (at(a+x))"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   950
      using s apply blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   951
     using a apply (auto simp: algebra_simps vd1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   952
     apply (force simp: shiftpath_def add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   953
    using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   954
    apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   955
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   956
  moreover
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   957
  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   958
        has_integral  integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {1 - a..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   959
    apply (rule has_integral_spike_finite
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   960
             [where s = "{1-a} \<union> (\<lambda>x. x-a+1) ` s" and f = "\<lambda>x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   961
      using s apply blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   962
     using a apply (auto simp: algebra_simps vd2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   963
     apply (force simp: shiftpath_def add.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   964
    using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   965
    apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   966
    apply (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   967
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   968
  ultimately show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   969
    using a
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   970
    by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   971
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   972
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   973
lemma has_contour_integral_shiftpath_D:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   974
  assumes "(f has_contour_integral i) (shiftpath a g)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   975
          "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   976
    shows "(f has_contour_integral i) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   977
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   978
  obtain s
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   979
    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
   980
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   981
  { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   982
    assume x: "0 < x" "x < 1" "x \<notin> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   983
    then have gx: "g differentiable at x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   984
      using g by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   985
    have "vector_derivative g (at x within {0..1}) =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   986
          vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   987
      apply (rule vector_derivative_at_within_ivl
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   988
                  [OF has_vector_derivative_transform_within_open
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   989
                      [of "{0<..<1}-s" _ "(shiftpath (1 - a) (shiftpath a g))"]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   990
      using s g assms x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   991
      apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   992
                        vector_derivative_within_interior vector_derivative_works [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   993
      apply (rule Derivative.differentiable_transform_at [of "min x (1-x)", OF _ _ gx])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   994
      apply (auto simp: dist_real_def shiftpath_shiftpath abs_if)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   995
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   996
  } note vd = this
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   997
  have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   998
    using assms  by (auto intro!: has_contour_integral_shiftpath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   999
  show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1000
    apply (simp add: has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1001
    apply (rule has_integral_spike_finite [of "{0,1} \<union> s", OF _ _  fi [unfolded has_contour_integral_def]])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1002
    using s assms vd
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1003
    apply (auto simp: Path_Connected.shiftpath_shiftpath)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1004
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1005
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1006
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1007
lemma has_contour_integral_shiftpath_eq:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1008
  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1009
    shows "(f has_contour_integral i) (shiftpath a g) \<longleftrightarrow> (f has_contour_integral i) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1010
  using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1011
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1012
lemma contour_integral_shiftpath:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1013
  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1014
    shows "contour_integral (shiftpath a g) f = contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1015
   using assms by (simp add: contour_integral_def has_contour_integral_shiftpath_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1016
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1017
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1018
subsection\<open>More about straight-line paths\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1019
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1020
lemma has_vector_derivative_linepath_within:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1021
    "(linepath a b has_vector_derivative (b - a)) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1022
apply (simp add: linepath_def has_vector_derivative_def algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1023
apply (rule derivative_eq_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1024
done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1025
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1026
lemma vector_derivative_linepath_within:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1027
    "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1028
  apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1029
  apply (auto simp: has_vector_derivative_linepath_within)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1030
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1031
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1032
lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1033
  by (simp add: has_vector_derivative_linepath_within vector_derivative_at)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1034
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1035
lemma valid_path_linepath [iff]: "valid_path (linepath a b)"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1036
  apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1037
  apply (rule_tac x="{}" in exI)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1038
  apply (simp add: differentiable_on_def differentiable_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1039
  using has_vector_derivative_def has_vector_derivative_linepath_within
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1040
  apply (fastforce simp add: continuous_on_eq_continuous_within)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1041
  done
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1042
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1043
lemma has_contour_integral_linepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1044
  shows "(f has_contour_integral i) (linepath a b) \<longleftrightarrow>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1045
         ((\<lambda>x. f(linepath a b x) * (b - a)) has_integral i) {0..1}"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1046
  by (simp add: has_contour_integral vector_derivative_linepath_at)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1047
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1048
lemma linepath_in_path:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1049
  shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1050
  by (auto simp: segment linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1051
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1052
lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1053
  by (auto simp: segment linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1054
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1055
lemma linepath_in_convex_hull:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1056
    fixes x::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1057
    assumes a: "a \<in> convex hull s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1058
        and b: "b \<in> convex hull s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1059
        and x: "0\<le>x" "x\<le>1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1060
       shows "linepath a b x \<in> convex hull s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1061
  apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1062
  using x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1063
  apply (auto simp: linepath_image_01 [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1064
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1065
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1066
lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1067
  by (simp add: linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1068
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1069
lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1070
  by (simp add: linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1071
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1072
lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1073
  by (simp add: scaleR_conv_of_real linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1074
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1075
lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1076
  by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1077
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1078
lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1079
  by (simp add: has_contour_integral_linepath)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1080
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1081
lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1082
  using has_contour_integral_trivial contour_integral_unique by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1083
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1084
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1085
subsection\<open>Relation to subpath construction\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1086
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1087
lemma valid_path_subpath:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1088
  fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1089
  assumes "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1090
    shows "valid_path(subpath u v g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1091
proof (cases "v=u")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1092
  case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1093
  then show ?thesis
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1094
    unfolding valid_path_def subpath_def
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1095
    by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1096
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1097
  case False
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1098
  have "(g o (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1099
    apply (rule piecewise_C1_differentiable_compose)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1100
    apply (simp add: C1_differentiable_imp_piecewise)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1101
     apply (simp add: image_affinity_atLeastAtMost)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1102
    using assms False
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1103
    apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1104
    apply (subst Int_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1105
    apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1106
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1107
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1108
    by (auto simp: o_def valid_path_def subpath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1109
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1110
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1111
lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1112
  by (simp add: has_contour_integral subpath_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1113
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1114
lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1115
  using has_contour_integral_subpath_refl contour_integrable_on_def by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1116
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1117
lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1118
  by (simp add: has_contour_integral_subpath_refl contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1119
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1120
lemma has_contour_integral_subpath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1121
  assumes f: "f contour_integrable_on g" and g: "valid_path g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1122
      and uv: "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1123
    shows "(f has_contour_integral  integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x)))
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1124
           (subpath u v g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1125
proof (cases "v=u")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1126
  case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1127
  then show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1128
    using f   by (simp add: contour_integrable_on_def subpath_def has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1129
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1130
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1131
  obtain s where s: "\<And>x. x \<in> {0..1} - s \<Longrightarrow> g differentiable at x" and fs: "finite s"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1132
    using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1133
  have *: "((\<lambda>x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u)))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1134
            has_integral (1 / (v - u)) * integral {u..v} (\<lambda>t. f (g t) * vector_derivative g (at t)))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1135
           {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1136
    using f uv
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1137
    apply (simp add: contour_integrable_on subpath_def has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1138
    apply (drule integrable_on_subcbox [where a=u and b=v, simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1139
    apply (simp_all add: has_integral_integral)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1140
    apply (drule has_integral_affinity [where m="v-u" and c=u, simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1141
    apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1142
    apply (simp add: divide_simps False)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1143
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1144
  { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1145
    have "x \<in> {0..1} \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1146
           x \<notin> (\<lambda>t. (v-u) *\<^sub>R t + u) -` s \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1147
           vector_derivative (\<lambda>x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1148
      apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1149
      apply (intro derivative_eq_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1150
      apply (cut_tac s [of "(v - u) * x + u"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1151
      using uv mult_left_le [of x "v-u"]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1152
      apply (auto simp:  vector_derivative_works)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1153
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1154
  } note vd = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1155
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1156
    apply (cut_tac has_integral_cmul [OF *, where c = "v-u"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1157
    using fs assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1158
    apply (simp add: False subpath_def has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1159
    apply (rule_tac s = "(\<lambda>t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1160
    apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1161
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1162
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1163
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1164
lemma contour_integrable_subpath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1165
  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1166
    shows "f contour_integrable_on (subpath u v g)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1167
  apply (cases u v rule: linorder_class.le_cases)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1168
   apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1169
  apply (subst reversepath_subpath [symmetric])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1170
  apply (rule contour_integrable_reversepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1171
   using assms apply (blast intro: valid_path_subpath)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1172
  apply (simp add: contour_integrable_on_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1173
  using assms apply (blast intro: has_contour_integral_subpath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1174
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1175
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1176
lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1177
  by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1178
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1179
lemma has_integral_contour_integral_subpath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1180
  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1181
    shows "(((\<lambda>x. f(g x) * vector_derivative g (at x)))
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1182
            has_integral  contour_integral (subpath u v g) f) {u..v}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1183
  using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1184
  apply (auto simp: has_integral_integrable_integral)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1185
  apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1186
  apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1187
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1188
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1189
lemma contour_integral_subcontour_integral:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1190
  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1191
    shows "contour_integral (subpath u v g) f =
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1192
           integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1193
  using assms has_contour_integral_subpath contour_integral_unique by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1194
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1195
lemma contour_integral_subpath_combine_less:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1196
  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1197
          "u<v" "v<w"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1198
    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1199
           contour_integral (subpath u w g) f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1200
  using assms apply (auto simp: contour_integral_subcontour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1201
  apply (rule integral_combine, auto)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1202
  apply (rule integrable_on_subcbox [where a=u and b=w and s = "{0..1}", simplified])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1203
  apply (auto simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1204
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1205
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1206
lemma contour_integral_subpath_combine:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1207
  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1208
    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1209
           contour_integral (subpath u w g) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1210
proof (cases "u\<noteq>v \<and> v\<noteq>w \<and> u\<noteq>w")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1211
  case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1212
    have *: "subpath v u g = reversepath(subpath u v g) \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1213
             subpath w u g = reversepath(subpath u w g) \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1214
             subpath w v g = reversepath(subpath v w g)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1215
      by (auto simp: reversepath_subpath)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1216
    have "u < v \<and> v < w \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1217
          u < w \<and> w < v \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1218
          v < u \<and> u < w \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1219
          v < w \<and> w < u \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1220
          w < u \<and> u < v \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1221
          w < v \<and> v < u"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1222
      using True assms by linarith
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1223
    with assms show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1224
      using contour_integral_subpath_combine_less [of f g u v w]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1225
            contour_integral_subpath_combine_less [of f g u w v]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1226
            contour_integral_subpath_combine_less [of f g v u w]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1227
            contour_integral_subpath_combine_less [of f g v w u]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1228
            contour_integral_subpath_combine_less [of f g w u v]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1229
            contour_integral_subpath_combine_less [of f g w v u]
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1230
      apply simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1231
      apply (elim disjE)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1232
      apply (auto simp: * contour_integral_reversepath contour_integrable_subpath
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1233
                   valid_path_reversepath valid_path_subpath algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1234
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1235
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1236
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1237
  then show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1238
    apply (auto simp: contour_integral_subpath_refl)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1239
    using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1240
    by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1241
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1242
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1243
lemma contour_integral_integral:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1244
  shows "contour_integral g f = integral {0..1} (\<lambda>x. f (g x) * vector_derivative g (at x))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1245
  by (simp add: contour_integral_def integral_def has_contour_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1246
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1247
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1248
subsection\<open>Segments via convex hulls\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1249
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1250
lemma segments_subset_convex_hull:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1251
    "closed_segment a b \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1252
    "closed_segment a c \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1253
    "closed_segment b c \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1254
    "closed_segment b a \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1255
    "closed_segment c a \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1256
    "closed_segment c b \<subseteq> (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1257
by (auto simp: segment_convex_hull linepath_of_real  elim!: rev_subsetD [OF _ hull_mono])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1258
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1259
lemma midpoints_in_convex_hull:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1260
  assumes "x \<in> convex hull s" "y \<in> convex hull s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1261
    shows "midpoint x y \<in> convex hull s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1262
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1263
  have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \<in> convex hull s"
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
  1264
    apply (rule convexD_alt)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1265
    using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1266
    apply (auto simp: convex_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1267
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1268
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1269
    by (simp add: midpoint_def algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1270
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1271
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1272
lemma convex_hull_subset:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1273
    "s \<subseteq> convex hull t \<Longrightarrow> convex hull s \<subseteq> convex hull t"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1274
  by (simp add: convex_convex_hull subset_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1275
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1276
lemma not_in_interior_convex_hull_3:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1277
  fixes a :: "complex"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1278
  shows "a \<notin> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1279
        "b \<notin> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1280
        "c \<notin> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1281
  by (auto simp: card_insert_le_m1 not_in_interior_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1282
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1283
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1284
text\<open>Cauchy's theorem where there's a primitive\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1285
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1286
lemma contour_integral_primitive_lemma:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1287
  fixes f :: "complex \<Rightarrow> complex" and g :: "real \<Rightarrow> complex"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1288
  assumes "a \<le> b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1289
      and "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1290
      and "g piecewise_differentiable_on {a..b}"  "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1291
    shows "((\<lambda>x. f'(g x) * vector_derivative g (at x within {a..b}))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1292
             has_integral (f(g b) - f(g a))) {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1293
proof -
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1294
  obtain k where k: "finite k" "\<forall>x\<in>{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1295
    using assms by (auto simp: piecewise_differentiable_on_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1296
  have cfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1297
    apply (rule continuous_on_compose [OF cg, unfolded o_def])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1298
    using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1299
    apply (metis complex_differentiable_def complex_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1300
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1301
  { fix x::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1302
    assume a: "a < x" and b: "x < b" and xk: "x \<notin> k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1303
    then have "g differentiable at x within {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1304
      using k by (simp add: differentiable_at_withinI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1305
    then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1306
      by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1307
    then have gdiff: "(g has_derivative (\<lambda>u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1308
      by (simp add: has_vector_derivative_def scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1309
    have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1310
      using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1311
    then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1312
      by (simp add: has_field_derivative_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1313
    have "((\<lambda>x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1314
      using diff_chain_within [OF gdiff fdiff]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1315
      by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1316
  } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1317
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1318
    apply (rule fundamental_theorem_of_calculus_interior_strong)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1319
    using k assms cfg *
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1320
    apply (auto simp: at_within_closed_interval)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1321
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1322
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1323
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1324
lemma contour_integral_primitive:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1325
  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1326
      and "valid_path g" "path_image g \<subseteq> s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1327
    shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1328
  using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1329
  apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1330
  apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1331
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1332
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1333
corollary Cauchy_theorem_primitive:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1334
  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1335
      and "valid_path g"  "path_image g \<subseteq> s" "pathfinish g = pathstart g"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1336
    shows "(f' has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1337
  using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1338
  by (metis diff_self contour_integral_primitive)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1339
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1340
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1341
text\<open>Existence of path integral for continuous function\<close>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1342
lemma contour_integrable_continuous_linepath:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1343
  assumes "continuous_on (closed_segment a b) f"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1344
  shows "f contour_integrable_on (linepath a b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1345
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1346
  have "continuous_on {0..1} ((\<lambda>x. f x * (b - a)) o linepath a b)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1347
    apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1348
    apply (rule continuous_intros | simp add: assms)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1349
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1350
  then show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1351
    apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1352
    apply (rule integrable_continuous [of 0 "1::real", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1353
    apply (rule continuous_on_eq [where f = "\<lambda>x. f(linepath a b x)*(b - a)"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1354
    apply (auto simp: vector_derivative_linepath_within)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1355
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1356
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1357
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1358
lemma has_field_der_id: "((\<lambda>x. x\<^sup>2 / 2) has_field_derivative x) (at x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1359
  by (rule has_derivative_imp_has_field_derivative)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1360
     (rule derivative_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1361
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1362
lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\<lambda>y. y) = (b^2 - a^2)/2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1363
  apply (rule contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1364
  using contour_integral_primitive [of UNIV "\<lambda>x. x^2/2" "\<lambda>x. x" "linepath a b"]
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1365
  apply (auto simp: field_simps has_field_der_id)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1366
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1367
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1368
lemma contour_integrable_on_const [iff]: "(\<lambda>x. c) contour_integrable_on (linepath a b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1369
  by (simp add: continuous_on_const contour_integrable_continuous_linepath)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1370
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1371
lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1372
  by (simp add: continuous_on_id contour_integrable_continuous_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1373
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1374
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1375
subsection\<open>Arithmetical combining theorems\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1376
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1377
lemma has_contour_integral_neg:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1378
    "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1379
  by (simp add: has_integral_neg has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1380
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1381
lemma has_contour_integral_add:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1382
    "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1383
     \<Longrightarrow> ((\<lambda>x. f1 x + f2 x) has_contour_integral (i1 + i2)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1384
  by (simp add: has_integral_add has_contour_integral_def algebra_simps)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1385
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1386
lemma has_contour_integral_diff:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1387
  "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1388
         \<Longrightarrow> ((\<lambda>x. f1 x - f2 x) has_contour_integral (i1 - i2)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1389
  by (simp add: has_integral_sub has_contour_integral_def algebra_simps)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1390
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1391
lemma has_contour_integral_lmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1392
  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. c * (f x)) has_contour_integral (c*i)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1393
apply (simp add: has_contour_integral_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1394
apply (drule has_integral_mult_right)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1395
apply (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1396
done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1397
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1398
lemma has_contour_integral_rmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1399
  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. (f x) * c) has_contour_integral (i*c)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1400
apply (drule has_contour_integral_lmul)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1401
apply (simp add: mult.commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1402
done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1403
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1404
lemma has_contour_integral_div:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1405
  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. f x/c) has_contour_integral (i/c)) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1406
  by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1407
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1408
lemma has_contour_integral_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1409
    "\<lbrakk>(f has_contour_integral y) p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> (g has_contour_integral y) p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1410
apply (simp add: path_image_def has_contour_integral_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1411
by (metis (no_types, lifting) image_eqI has_integral_eq)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1412
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1413
lemma has_contour_integral_bound_linepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1414
  assumes "(f has_contour_integral i) (linepath a b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1415
          "0 \<le> B" "\<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1416
    shows "norm i \<le> B * norm(b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1417
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1418
  { fix x::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1419
    assume x: "0 \<le> x" "x \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1420
  have "norm (f (linepath a b x)) *
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1421
        norm (vector_derivative (linepath a b) (at x within {0..1})) \<le> B * norm (b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1422
    by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1423
  } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1424
  have "norm i \<le> (B * norm (b - a)) * content (cbox 0 (1::real))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1425
    apply (rule has_integral_bound
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1426
       [of _ "\<lambda>x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1427
    using assms * unfolding has_contour_integral_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1428
    apply (auto simp: norm_mult)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1429
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1430
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1431
    by (auto simp: content_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1432
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1433
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1434
(*UNUSED
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1435
lemma has_contour_integral_bound_linepath_strong:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1436
  fixes a :: real and f :: "complex \<Rightarrow> real"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1437
  assumes "(f has_contour_integral i) (linepath a b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1438
          "finite k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1439
          "0 \<le> B" "\<And>x::real. x \<in> closed_segment a b - k \<Longrightarrow> norm(f x) \<le> B"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1440
    shows "norm i \<le> B*norm(b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1441
*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1442
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1443
lemma has_contour_integral_const_linepath: "((\<lambda>x. c) has_contour_integral c*(b - a))(linepath a b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1444
  unfolding has_contour_integral_linepath
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1445
  by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1446
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1447
lemma has_contour_integral_0: "((\<lambda>x. 0) has_contour_integral 0) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1448
  by (simp add: has_contour_integral_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1449
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1450
lemma has_contour_integral_is_0:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1451
    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1452
  by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1453
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1454
lemma has_contour_integral_setsum:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1455
    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1456
     \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1457
  by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1458
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1459
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1460
subsection \<open>Operations on path integrals\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1461
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1462
lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1463
  by (rule contour_integral_unique [OF has_contour_integral_const_linepath])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1464
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1465
lemma contour_integral_neg:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1466
    "f contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. -(f x)) = -(contour_integral g f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1467
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1468
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1469
lemma contour_integral_add:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1470
    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x + f2 x) =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1471
                contour_integral g f1 + contour_integral g f2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1472
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1473
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1474
lemma contour_integral_diff:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1475
    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x - f2 x) =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1476
                contour_integral g f1 - contour_integral g f2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1477
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1478
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1479
lemma contour_integral_lmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1480
  shows "f contour_integrable_on g
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1481
           \<Longrightarrow> contour_integral g (\<lambda>x. c * f x) = c*contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1482
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1483
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1484
lemma contour_integral_rmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1485
  shows "f contour_integrable_on g
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1486
        \<Longrightarrow> contour_integral g (\<lambda>x. f x * c) = contour_integral g f * c"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1487
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1488
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1489
lemma contour_integral_div:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1490
  shows "f contour_integrable_on g
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1491
        \<Longrightarrow> contour_integral g (\<lambda>x. f x / c) = contour_integral g f / c"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1492
  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1493
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1494
lemma contour_integral_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1495
    "(\<And>x. x \<in> path_image p \<Longrightarrow> f x = g x) \<Longrightarrow> contour_integral p f = contour_integral p g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1496
  by (simp add: contour_integral_def) (metis has_contour_integral_eq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1497
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1498
lemma contour_integral_eq_0:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1499
    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> contour_integral g f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1500
  by (simp add: has_contour_integral_is_0 contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1501
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1502
lemma contour_integral_bound_linepath:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1503
  shows
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1504
    "\<lbrakk>f contour_integrable_on (linepath a b);
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1505
      0 \<le> B; \<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1506
     \<Longrightarrow> norm(contour_integral (linepath a b) f) \<le> B*norm(b - a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1507
  apply (rule has_contour_integral_bound_linepath [of f])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1508
  apply (auto simp: has_contour_integral_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1509
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1510
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1511
lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1512
  by (simp add: contour_integral_unique has_contour_integral_0)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1513
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1514
lemma contour_integral_setsum:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1515
    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1516
     \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1517
  by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1518
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1519
lemma contour_integrable_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1520
    "\<lbrakk>f contour_integrable_on p; \<And>x. x \<in> path_image p \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g contour_integrable_on p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1521
  unfolding contour_integrable_on_def
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1522
  by (metis has_contour_integral_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1523
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1524
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1525
subsection \<open>Arithmetic theorems for path integrability\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1526
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1527
lemma contour_integrable_neg:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1528
    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1529
  using has_contour_integral_neg contour_integrable_on_def by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1530
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1531
lemma contour_integrable_add:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1532
    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x + f2 x) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1533
  using has_contour_integral_add contour_integrable_on_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1534
  by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1535
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1536
lemma contour_integrable_diff:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1537
    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x - f2 x) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1538
  using has_contour_integral_diff contour_integrable_on_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1539
  by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1540
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1541
lemma contour_integrable_lmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1542
    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. c * f x) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1543
  using has_contour_integral_lmul contour_integrable_on_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1544
  by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1545
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1546
lemma contour_integrable_rmul:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1547
    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x * c) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1548
  using has_contour_integral_rmul contour_integrable_on_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1549
  by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1550
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1551
lemma contour_integrable_div:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1552
    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x / c) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1553
  using has_contour_integral_div contour_integrable_on_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1554
  by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1555
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1556
lemma contour_integrable_setsum:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1557
    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1558
     \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1559
   unfolding contour_integrable_on_def
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1560
   by (metis has_contour_integral_setsum)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1561
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1562
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1563
subsection\<open>Reversing a path integral\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1564
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1565
lemma has_contour_integral_reverse_linepath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1566
    "(f has_contour_integral i) (linepath a b)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1567
     \<Longrightarrow> (f has_contour_integral (-i)) (linepath b a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1568
  using has_contour_integral_reversepath valid_path_linepath by fastforce
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1569
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1570
lemma contour_integral_reverse_linepath:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1571
    "continuous_on (closed_segment a b) f
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1572
     \<Longrightarrow> contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1573
apply (rule contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1574
apply (rule has_contour_integral_reverse_linepath)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1575
by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1576
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1577
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1578
(* Splitting a path integral in a flat way.*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1579
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1580
lemma has_contour_integral_split:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1581
  assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1582
      and k: "0 \<le> k" "k \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1583
      and c: "c - a = k *\<^sub>R (b - a)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1584
    shows "(f has_contour_integral (i + j)) (linepath a b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1585
proof (cases "k = 0 \<or> k = 1")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1586
  case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1587
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1588
    using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1589
    apply auto
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1590
    apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1591
    apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1592
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1593
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1594
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1595
  then have k: "0 < k" "k < 1" "complex_of_real k \<noteq> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1596
    using assms apply auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1597
    using of_real_eq_iff by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1598
  have c': "c = k *\<^sub>R (b - a) + a"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1599
    by (metis diff_add_cancel c)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1600
  have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1601
    by (simp add: algebra_simps c')
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1602
  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1603
    have **: "\<And>x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1604
      using False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1605
      apply (simp add: c' algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1606
      apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1607
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1608
    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1609
      using * k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1610
      apply -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1611
      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1612
      apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1613
      apply (drule Integration.has_integral_cmul [where c = "inverse k"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1614
      apply (simp add: Integration.has_integral_cmul)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1615
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1616
  } note fi = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1617
  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1618
    have **: "\<And>x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1619
      using k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1620
      apply (simp add: c' field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1621
      apply (simp add: scaleR_conv_of_real divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1622
      apply (simp add: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1623
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1624
    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1625
      using * k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1626
      apply -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1627
      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1628
      apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1629
      apply (drule Integration.has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1630
      apply (simp add: Integration.has_integral_cmul)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1631
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1632
  } note fj = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1633
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1634
    using f k
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1635
    apply (simp add: has_contour_integral_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1636
    apply (simp add: linepath_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1637
    apply (rule has_integral_combine [OF _ _ fi fj], simp_all)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1638
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1639
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1640
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1641
lemma continuous_on_closed_segment_transform:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1642
  assumes f: "continuous_on (closed_segment a b) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1643
      and k: "0 \<le> k" "k \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1644
      and c: "c - a = k *\<^sub>R (b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1645
    shows "continuous_on (closed_segment a c) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1646
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1647
  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1648
    using c by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1649
  show "continuous_on (closed_segment a c) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1650
    apply (rule continuous_on_subset [OF f])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1651
    apply (simp add: segment_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1652
    apply (rule convex_hull_subset)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1653
    using assms
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
  1654
    apply (auto simp: hull_inc c' Convex.convexD_alt)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1655
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1656
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1657
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1658
lemma contour_integral_split:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1659
  assumes f: "continuous_on (closed_segment a b) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1660
      and k: "0 \<le> k" "k \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1661
      and c: "c - a = k *\<^sub>R (b - a)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1662
    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1663
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1664
  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1665
    using c by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1666
  have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1667
    apply (rule_tac [!] continuous_on_subset [OF f])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1668
    apply (simp_all add: segment_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1669
    apply (rule_tac [!] convex_hull_subset)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1670
    using assms
61426
d53db136e8fd new material on path_component_sets, inside, outside, etc. And more default simprules
paulson <lp15@cam.ac.uk>
parents: 61284
diff changeset
  1671
    apply (auto simp: hull_inc c' Convex.convexD_alt)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1672
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1673
  show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1674
    apply (rule contour_integral_unique)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1675
    apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1676
    apply (rule contour_integrable_continuous_linepath *)+
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1677
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1678
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1679
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1680
lemma contour_integral_split_linepath:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1681
  assumes f: "continuous_on (closed_segment a b) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1682
      and c: "c \<in> closed_segment a b"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1683
    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1684
  using c
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1685
  by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1686
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1687
(* The special case of midpoints used in the main quadrisection.*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1688
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1689
lemma has_contour_integral_midpoint:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1690
  assumes "(f has_contour_integral i) (linepath a (midpoint a b))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1691
          "(f has_contour_integral j) (linepath (midpoint a b) b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1692
    shows "(f has_contour_integral (i + j)) (linepath a b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1693
  apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1694
  using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1695
  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1696
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1697
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1698
lemma contour_integral_midpoint:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1699
   "continuous_on (closed_segment a b) f
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1700
    \<Longrightarrow> contour_integral (linepath a b) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1701
        contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1702
  apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1703
  using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1704
  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1705
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1706
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1707
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1708
text\<open>A couple of special case lemmas that are useful below\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1709
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1710
lemma triangle_linear_has_chain_integral:
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1711
    "((\<lambda>x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1712
  apply (rule Cauchy_theorem_primitive [of UNIV "\<lambda>x. m/2 * x^2 + d*x"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1713
  apply (auto intro!: derivative_eq_intros)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1714
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1715
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1716
lemma has_chain_integral_chain_integral3:
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1717
     "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1718
      \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1719
  apply (subst contour_integral_unique [symmetric], assumption)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1720
  apply (drule has_contour_integral_integrable)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1721
  apply (simp add: valid_path_join)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1722
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1723
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1724
subsection\<open>Reversing the order in a double path integral\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1725
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1726
text\<open>The condition is stronger than needed but it's often true in typical situations\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1727
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1728
lemma fst_im_cbox [simp]: "cbox c d \<noteq> {} \<Longrightarrow> (fst ` cbox (a,c) (b,d)) = cbox a b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1729
  by (auto simp: cbox_Pair_eq)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1730
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1731
lemma snd_im_cbox [simp]: "cbox a b \<noteq> {} \<Longrightarrow> (snd ` cbox (a,c) (b,d)) = cbox c d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1732
  by (auto simp: cbox_Pair_eq)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1733
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1734
lemma contour_integral_swap:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1735
  assumes fcon:  "continuous_on (path_image g \<times> path_image h) (\<lambda>(y1,y2). f y1 y2)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1736
      and vp:    "valid_path g" "valid_path h"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1737
      and gvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative g (at t))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1738
      and hvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative h (at t))"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1739
  shows "contour_integral g (\<lambda>w. contour_integral h (f w)) =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1740
         contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1741
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1742
  have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  1743
    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1744
  have fgh1: "\<And>x. (\<lambda>t. f (g x) (h t)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g x, h t))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1745
    by (rule ext) simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1746
  have fgh2: "\<And>x. (\<lambda>t. f (g t) (h x)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g t, h x))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1747
    by (rule ext) simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1748
  have fcon_im1: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g x, h t)) ` {0..1}) (\<lambda>(x, y). f x y)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1749
    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1750
  have fcon_im2: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> continuous_on ((\<lambda>t. (g t, h x)) ` {0..1}) (\<lambda>(x, y). f x y)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1751
    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1752
  have vdg: "\<And>y. y \<in> {0..1} \<Longrightarrow> (\<lambda>x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1753
    apply (rule integrable_continuous_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1754
    apply (rule continuous_on_mult [OF _ gvcon])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1755
    apply (subst fgh2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1756
    apply (rule fcon_im2 gcon continuous_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1757
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1758
  have "(\<lambda>z. vector_derivative g (at (fst z))) = (\<lambda>x. vector_derivative g (at x)) o fst"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1759
    by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1760
  then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>x. vector_derivative g (at (fst x)))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1761
    apply (rule ssubst)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1762
    apply (rule continuous_intros | simp add: gvcon)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1763
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1764
  have "(\<lambda>z. vector_derivative h (at (snd z))) = (\<lambda>x. vector_derivative h (at x)) o snd"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1765
    by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1766
  then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1767
    apply (rule ssubst)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1768
    apply (rule continuous_intros | simp add: hvcon)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1769
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1770
  have "(\<lambda>x. f (g (fst x)) (h (snd x))) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>w. ((g o fst) w, (h o snd) w))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1771
    by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1772
  then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1773
    apply (rule ssubst)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1774
    apply (rule gcon hcon continuous_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1775
    apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1776
    done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1777
  have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1778
        integral {0..1} (\<lambda>x. contour_integral h (\<lambda>y. f (g x) y * vector_derivative g (at x)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1779
    apply (rule integral_cong [OF contour_integral_rmul [symmetric]])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1780
    apply (clarsimp simp: contour_integrable_on)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1781
    apply (rule integrable_continuous_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1782
    apply (rule continuous_on_mult [OF _ hvcon])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1783
    apply (subst fgh1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1784
    apply (rule fcon_im1 hcon continuous_intros | simp)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1785
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1786
  also have "... = integral {0..1}
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1787
                     (\<lambda>y. contour_integral g (\<lambda>x. f x (h y) * vector_derivative h (at y)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1788
    apply (simp add: contour_integral_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1789
    apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1790
    apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1791
    apply (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1792
    done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1793
  also have "... = contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1794
    apply (simp add: contour_integral_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1795
    apply (rule integral_cong)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1796
    apply (subst integral_mult_left [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1797
    apply (blast intro: vdg)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1798
    apply (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1799
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1800
  finally show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1801
    by (simp add: contour_integral_integral)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1802
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1803
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1804
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1805
subsection\<open>The key quadrisection step\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1806
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1807
lemma norm_sum_half:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1808
  assumes "norm(a + b) >= e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1809
    shows "norm a >= e/2 \<or> norm b >= e/2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1810
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1811
  have "e \<le> norm (- a - b)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1812
    by (simp add: add.commute assms norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1813
  thus ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1814
    using norm_triangle_ineq4 order_trans by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1815
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1816
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1817
lemma norm_sum_lemma:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1818
  assumes "e \<le> norm (a + b + c + d)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1819
    shows "e / 4 \<le> norm a \<or> e / 4 \<le> norm b \<or> e / 4 \<le> norm c \<or> e / 4 \<le> norm d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1820
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1821
  have "e \<le> norm ((a + b) + (c + d))" using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1822
    by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1823
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1824
    by (auto dest!: norm_sum_half)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1825
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1826
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1827
lemma Cauchy_theorem_quadrisection:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1828
  assumes f: "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1829
      and dist: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1830
      and e: "e * K^2 \<le>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1831
              norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1832
  shows "\<exists>a' b' c'.
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1833
           a' \<in> convex hull {a,b,c} \<and> b' \<in> convex hull {a,b,c} \<and> c' \<in> convex hull {a,b,c} \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1834
           dist a' b' \<le> K/2  \<and>  dist b' c' \<le> K/2  \<and>  dist c' a' \<le> K/2  \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1835
           e * (K/2)^2 \<le> norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1836
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1837
  note divide_le_eq_numeral1 [simp del]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1838
  def a' \<equiv> "midpoint b c"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1839
  def b' \<equiv> "midpoint c a"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1840
  def c' \<equiv> "midpoint a b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1841
  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1842
    using f continuous_on_subset segments_subset_convex_hull by metis+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1843
  have fcont': "continuous_on (closed_segment c' b') f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1844
               "continuous_on (closed_segment a' c') f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1845
               "continuous_on (closed_segment b' a') f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1846
    unfolding a'_def b'_def c'_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1847
    apply (rule continuous_on_subset [OF f],
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1848
           metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1849
    done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1850
  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1851
  have *: "?pathint a b + ?pathint b c + ?pathint c a =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1852
          (?pathint a c' + ?pathint c' b' + ?pathint b' a) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1853
          (?pathint a' c' + ?pathint c' b + ?pathint b a') +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1854
          (?pathint a' c + ?pathint c b' + ?pathint b' a') +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1855
          (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1856
    apply (simp add: fcont' contour_integral_reverse_linepath)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1857
    apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1858
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1859
  have [simp]: "\<And>x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1860
    by (metis left_diff_distrib mult.commute norm_mult_numeral1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1861
  have [simp]: "\<And>x y. cmod (x - y) = cmod (y - x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1862
    by (simp add: norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1863
  consider "e * K\<^sup>2 / 4 \<le> cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" |
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1864
           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" |
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1865
           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" |
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1866
           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1867
    using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1868
    apply (simp only: *)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1869
    apply (blast intro: that dest!: norm_sum_lemma)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1870
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1871
  then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1872
  proof cases
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1873
    case 1 then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1874
      apply (rule_tac x=a in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1875
      apply (rule exI [where x=c'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1876
      apply (rule exI [where x=b'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1877
      using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1878
      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1879
      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1880
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1881
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1882
    case 2 then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1883
      apply (rule_tac x=a' in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1884
      apply (rule exI [where x=c'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1885
      apply (rule exI [where x=b])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1886
      using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1887
      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1888
      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1889
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1890
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1891
    case 3 then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1892
      apply (rule_tac x=a' in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1893
      apply (rule exI [where x=c])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1894
      apply (rule exI [where x=b'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1895
      using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1896
      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1897
      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1898
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1899
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1900
    case 4 then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1901
      apply (rule_tac x=a' in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1902
      apply (rule exI [where x=b'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1903
      apply (rule exI [where x=c'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1904
      using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1905
      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1906
      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1907
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1908
  qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1909
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1910
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1911
subsection\<open>Cauchy's theorem for triangles\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1912
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1913
lemma triangle_points_closer:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1914
  fixes a::complex
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1915
  shows "\<lbrakk>x \<in> convex hull {a,b,c};  y \<in> convex hull {a,b,c}\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1916
         \<Longrightarrow> norm(x - y) \<le> norm(a - b) \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1917
             norm(x - y) \<le> norm(b - c) \<or>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1918
             norm(x - y) \<le> norm(c - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1919
  using simplex_extremal_le [of "{a,b,c}"]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1920
  by (auto simp: norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1921
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1922
lemma holomorphic_point_small_triangle:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1923
  assumes x: "x \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1924
      and f: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1925
      and cd: "f complex_differentiable (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1926
      and e: "0 < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1927
    shows "\<exists>k>0. \<forall>a b c. dist a b \<le> k \<and> dist b c \<le> k \<and> dist c a \<le> k \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1928
              x \<in> convex hull {a,b,c} \<and> convex hull {a,b,c} \<subseteq> s
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1929
              \<longrightarrow> norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1930
                       contour_integral(linepath c a) f)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1931
                  \<le> e*(dist a b + dist b c + dist c a)^2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1932
           (is "\<exists>k>0. \<forall>a b c. _ \<longrightarrow> ?normle a b c")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1933
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1934
  have le_of_3: "\<And>a x y z. \<lbrakk>0 \<le> x*y; 0 \<le> x*z; 0 \<le> y*z; a \<le> (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1935
                     \<Longrightarrow> a \<le> e*(x + y + z)^2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1936
    by (simp add: algebra_simps power2_eq_square)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1937
  have disj_le: "\<lbrakk>x \<le> a \<or> x \<le> b \<or> x \<le> c; 0 \<le> a; 0 \<le> b; 0 \<le> c\<rbrakk> \<Longrightarrow> x \<le> a + b + c"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1938
             for x::real and a b c
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1939
    by linarith
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1940
  have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1941
              if "convex hull {a, b, c} \<subseteq> s" for a b c
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1942
    using segments_subset_convex_hull that
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1943
    by (metis continuous_on_subset f contour_integrable_continuous_linepath)+
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1944
  note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral]
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1945
  { fix f' a b c d
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1946
    assume d: "0 < d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1947
       and f': "\<And>y. \<lbrakk>cmod (y - x) \<le> d; y \<in> s\<rbrakk> \<Longrightarrow> cmod (f y - f x - f' * (y - x)) \<le> e * cmod (y - x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1948
       and le: "cmod (a - b) \<le> d" "cmod (b - c) \<le> d" "cmod (c - a) \<le> d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1949
       and xc: "x \<in> convex hull {a, b, c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1950
       and s: "convex hull {a, b, c} \<subseteq> s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1951
    have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1952
              contour_integral (linepath a b) (\<lambda>y. f y - f x - f'*(y - x)) +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1953
              contour_integral (linepath b c) (\<lambda>y. f y - f x - f'*(y - x)) +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1954
              contour_integral (linepath c a) (\<lambda>y. f y - f x - f'*(y - x))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1955
      apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1956
      apply (simp add: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1957
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1958
    { fix y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1959
      assume yc: "y \<in> convex hull {a,b,c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1960
      have "cmod (f y - f x - f' * (y - x)) \<le> e*norm(y - x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1961
        apply (rule f')
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1962
        apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1963
        using s yc by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1964
      also have "... \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1965
        by (simp add: yc e xc disj_le [OF triangle_points_closer])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1966
      finally have "cmod (f y - f x - f' * (y - x)) \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" .
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1967
    } note cm_le = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1968
    have "?normle a b c"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1969
      apply (simp add: dist_norm pa)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1970
      apply (rule le_of_3)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1971
      using f' xc s e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1972
      apply simp_all
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1973
      apply (intro norm_triangle_le add_mono path_bound)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1974
      apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1975
      apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1976
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1977
  } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1978
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1979
    using cd e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1980
    apply (simp add: complex_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1981
    apply (clarify dest!: spec mp)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1982
    using *
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1983
    apply (simp add: dist_norm, blast)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1984
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1985
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1986
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1987
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1988
(* Hence the most basic theorem for a triangle.*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1989
locale Chain =
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1990
  fixes x0 At Follows
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1991
  assumes At0: "At x0 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1992
      and AtSuc: "\<And>x n. At x n \<Longrightarrow> \<exists>x'. At x' (Suc n) \<and> Follows x' x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1993
begin
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1994
  primrec f where
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1995
    "f 0 = x0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1996
  | "f (Suc n) = (SOME x. At x (Suc n) \<and> Follows x (f n))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1997
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1998
  lemma At: "At (f n) n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  1999
  proof (induct n)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2000
    case 0 show ?case
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2001
      by (simp add: At0)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2002
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2003
    case (Suc n) show ?case
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2004
      by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2005
  qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2006
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2007
  lemma Follows: "Follows (f(Suc n)) (f n)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2008
    by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2009
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2010
  declare f.simps(2) [simp del]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2011
end
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2012
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2013
lemma Chain3:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2014
  assumes At0: "At x0 y0 z0 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2015
      and AtSuc: "\<And>x y z n. At x y z n \<Longrightarrow> \<exists>x' y' z'. At x' y' z' (Suc n) \<and> Follows x' y' z' x y z"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2016
  obtains f g h where
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2017
    "f 0 = x0" "g 0 = y0" "h 0 = z0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2018
                      "\<And>n. At (f n) (g n) (h n) n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2019
                       "\<And>n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2020
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2021
  interpret three: Chain "(x0,y0,z0)" "\<lambda>(x,y,z). At x y z" "\<lambda>(x',y',z'). \<lambda>(x,y,z). Follows x' y' z' x y z"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2022
    apply unfold_locales
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2023
    using At0 AtSuc by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2024
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2025
  apply (rule that [of "\<lambda>n. fst (three.f n)"  "\<lambda>n. fst (snd (three.f n))" "\<lambda>n. snd (snd (three.f n))"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2026
  apply simp_all
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2027
  using three.At three.Follows
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2028
  apply (simp_all add: split_beta')
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2029
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2030
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2031
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2032
lemma Cauchy_theorem_triangle:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2033
  assumes "f holomorphic_on (convex hull {a,b,c})"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2034
    shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2035
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2036
  have contf: "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2037
    by (metis assms holomorphic_on_imp_continuous_on)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2038
  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2039
  { fix y::complex
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2040
    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2041
       and ynz: "y \<noteq> 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2042
    def K \<equiv> "1 + max (dist a b) (max (dist b c) (dist c a))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2043
    def e \<equiv> "norm y / K^2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2044
    have K1: "K \<ge> 1"  by (simp add: K_def max.coboundedI1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2045
    then have K: "K > 0" by linarith
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2046
    have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2047
      by (simp_all add: K_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2048
    have e: "e > 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2049
      unfolding e_def using ynz K1 by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2050
    def At \<equiv> "\<lambda>x y z n. convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2051
                         dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2052
                         norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2053
    have At0: "At a b c 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2054
      using fy
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2055
      by (simp add: At_def e_def has_chain_integral_chain_integral3)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2056
    { fix x y z n
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2057
      assume At: "At x y z n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2058
      then have contf': "continuous_on (convex hull {x,y,z}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2059
        using contf At_def continuous_on_subset by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2060
      have "\<exists>x' y' z'. At x' y' z' (Suc n) \<and> convex hull {x',y',z'} \<subseteq> convex hull {x,y,z}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2061
        using At
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2062
        apply (simp add: At_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2063
        using  Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2064
        apply clarsimp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2065
        apply (rule_tac x="a'" in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2066
        apply (rule_tac x="b'" in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2067
        apply (rule_tac x="c'" in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2068
        apply (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2069
        apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2070
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2071
    } note AtSuc = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2072
    obtain fa fb fc
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2073
      where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2074
        and cosb: "\<And>n. convex hull {fa n, fb n, fc n} \<subseteq> convex hull {a,b,c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2075
        and dist: "\<And>n. dist (fa n) (fb n) \<le> K/2^n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2076
                  "\<And>n. dist (fb n) (fc n) \<le> K/2^n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2077
                  "\<And>n. dist (fc n) (fa n) \<le> K/2^n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2078
        and no: "\<And>n. norm(?pathint (fa n) (fb n) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2079
                           ?pathint (fb n) (fc n) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2080
                           ?pathint (fc n) (fa n)) \<ge> e * (K/2^n)^2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2081
        and conv_le: "\<And>n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \<subseteq> convex hull {fa n, fb n, fc n}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2082
      apply (rule Chain3 [of At, OF At0 AtSuc])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2083
      apply (auto simp: At_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2084
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2085
    have "\<exists>x. \<forall>n. x \<in> convex hull {fa n, fb n, fc n}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2086
      apply (rule bounded_closed_nest)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2087
      apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2088
      apply (rule allI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2089
      apply (rule transitive_stepwise_le)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2090
      apply (auto simp: conv_le)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2091
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2092
    then obtain x where x: "\<And>n. x \<in> convex hull {fa n, fb n, fc n}" by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2093
    then have xin: "x \<in> convex hull {a,b,c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2094
      using assms f0 by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2095
    then have fx: "f complex_differentiable at x within (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2096
      using assms holomorphic_on_def by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2097
    { fix k n
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2098
      assume k: "0 < k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2099
         and le:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2100
            "\<And>x' y' z'.
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2101
               \<lbrakk>dist x' y' \<le> k; dist y' z' \<le> k; dist z' x' \<le> k;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2102
                x \<in> convex hull {x',y',z'};
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2103
                convex hull {x',y',z'} \<subseteq> convex hull {a,b,c}\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2104
               \<Longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2105
               cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2106
                     \<le> e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2107
         and Kk: "K / k < 2 ^ n"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2108
      have "K / 2 ^ n < k" using Kk k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2109
        by (auto simp: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2110
      then have DD: "dist (fa n) (fb n) \<le> k" "dist (fb n) (fc n) \<le> k" "dist (fc n) (fa n) \<le> k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2111
        using dist [of n]  k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2112
        by linarith+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2113
      have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2114
               \<le> (3 * K / 2 ^ n)\<^sup>2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2115
        using dist [of n] e K
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2116
        by (simp add: abs_le_square_iff [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2117
      have less10: "\<And>x y::real. 0 < x \<Longrightarrow> y \<le> 9*x \<Longrightarrow> y < x*10"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2118
        by linarith
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2119
      have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \<le> e * (3 * K / 2 ^ n)\<^sup>2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2120
        using ynz dle e mult_le_cancel_left_pos by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2121
      also have "... <
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2122
          cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2123
        using no [of n] e K
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2124
        apply (simp add: e_def field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2125
        apply (simp only: zero_less_norm_iff [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2126
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2127
      finally have False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2128
        using le [OF DD x cosb] by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2129
    } then
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2130
    have ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2131
      using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2132
      apply clarsimp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2133
      apply (rule_tac x1="K/k" in exE [OF real_arch_pow2], blast)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2134
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2135
  }
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2136
  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2137
    by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2138
                   segments_subset_convex_hull(3) segments_subset_convex_hull(5))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2139
  ultimately show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2140
    using has_contour_integral_integral by fastforce
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2141
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2142
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2143
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2144
subsection\<open>Version needing function holomorphic in interior only\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2145
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2146
lemma Cauchy_theorem_flat_lemma:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2147
  assumes f: "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2148
      and c: "c - a = k *\<^sub>R (b - a)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2149
      and k: "0 \<le> k"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2150
    shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2151
          contour_integral (linepath c a) f = 0"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2152
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2153
  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2154
    using f continuous_on_subset segments_subset_convex_hull by metis+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2155
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2156
  proof (cases "k \<le> 1")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2157
    case True show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2158
      by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2159
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2160
    case False then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2161
      using fabc c
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2162
      apply (subst contour_integral_split [of a c f "1/k" b, symmetric])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2163
      apply (metis closed_segment_commute fabc(3))
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2164
      apply (auto simp: k contour_integral_reverse_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2165
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2166
  qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2167
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2168
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2169
lemma Cauchy_theorem_flat:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2170
  assumes f: "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2171
      and c: "c - a = k *\<^sub>R (b - a)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2172
    shows "contour_integral (linepath a b) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2173
           contour_integral (linepath b c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2174
           contour_integral (linepath c a) f = 0"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2175
proof (cases "0 \<le> k")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2176
  case True with assms show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2177
    by (blast intro: Cauchy_theorem_flat_lemma)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2178
next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2179
  case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2180
  have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2181
    using f continuous_on_subset segments_subset_convex_hull by metis+
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2182
  moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2183
        contour_integral (linepath c b) f = 0"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2184
    apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2185
    using False c
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2186
    apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2187
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2188
  ultimately show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2189
    apply (auto simp: contour_integral_reverse_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2190
    using add_eq_0_iff by force
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2191
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2192
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2193
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2194
lemma Cauchy_theorem_triangle_interior:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2195
  assumes contf: "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2196
      and holf:  "f holomorphic_on interior (convex hull {a,b,c})"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2197
     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2198
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2199
  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2200
    using contf continuous_on_subset segments_subset_convex_hull by metis+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2201
  have "bounded (f ` (convex hull {a,b,c}))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2202
    by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2203
  then obtain B where "0 < B" and Bnf: "\<And>x. x \<in> convex hull {a,b,c} \<Longrightarrow> norm (f x) \<le> B"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2204
     by (auto simp: dest!: bounded_pos [THEN iffD1])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2205
  have "bounded (convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2206
    by (simp add: bounded_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2207
  then obtain C where C: "0 < C" and Cno: "\<And>y. y \<in> convex hull {a,b,c} \<Longrightarrow> norm y < C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2208
    using bounded_pos_less by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2209
  then have diff_2C: "norm(x - y) \<le> 2*C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2210
           if x: "x \<in> convex hull {a, b, c}" and y: "y \<in> convex hull {a, b, c}" for x y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2211
  proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2212
    have "cmod x \<le> C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2213
      using x by (meson Cno not_le not_less_iff_gr_or_eq)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2214
    hence "cmod (x - y) \<le> C + C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2215
      using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2216
    thus "cmod (x - y) \<le> 2 * C"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2217
      by (metis mult_2)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2218
  qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2219
  have contf': "continuous_on (convex hull {b,a,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2220
    using contf by (simp add: insert_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2221
  { fix y::complex
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2222
    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2223
       and ynz: "y \<noteq> 0"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2224
    have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2225
      by (rule has_chain_integral_chain_integral3 [OF fy])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2226
    have ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2227
    proof (cases "c=a \<or> a=b \<or> b=c")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2228
      case True then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2229
        using Cauchy_theorem_flat [OF contf, of 0]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2230
        using has_chain_integral_chain_integral3 [OF fy] ynz
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2231
        by (force simp: fabc contour_integral_reverse_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2232
    next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2233
      case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2234
      then have car3: "card {a, b, c} = Suc (DIM(complex))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2235
        by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2236
      { assume "interior(convex hull {a,b,c}) = {}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2237
        then have "collinear{a,b,c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2238
          using interior_convex_hull_eq_empty [OF car3]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2239
          by (simp add: collinear_3_eq_affine_dependent)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2240
        then have "False"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2241
          using False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2242
          apply (clarsimp simp add: collinear_3 collinear_lemma)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2243
          apply (drule Cauchy_theorem_flat [OF contf'])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2244
          using pi_eq_y ynz
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2245
          apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2246
          done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2247
      }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2248
      then obtain d where d: "d \<in> interior (convex hull {a, b, c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2249
        by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2250
      { fix d1
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2251
        assume d1_pos: "0 < d1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2252
           and d1: "\<And>x x'. \<lbrakk>x\<in>convex hull {a, b, c}; x'\<in>convex hull {a, b, c}; cmod (x' - x) < d1\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2253
                           \<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2254
        def e      \<equiv> "min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2255
        def shrink \<equiv> "\<lambda>x. x - e *\<^sub>R (x - d)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2256
        let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2257
        have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2258
          using d1_pos \<open>C>0\<close> \<open>B>0\<close> ynz by (simp_all add: e_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2259
        then have eCB: "24 * e * C * B \<le> cmod y"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2260
          using \<open>C>0\<close> \<open>B>0\<close>  by (simp add: field_simps)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2261
        have e_le_d1: "e * (4 * C) \<le> d1"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2262
          using e \<open>C>0\<close> by (simp add: field_simps)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2263
        have "shrink a \<in> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2264
             "shrink b \<in> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2265
             "shrink c \<in> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2266
          using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2267
        then have fhp0: "(f has_contour_integral 0)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2268
                (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2269
          by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2270
        then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2271
          by (simp add: has_chain_integral_chain_integral3)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2272
        have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2273
                      "f contour_integrable_on linepath (shrink b) (shrink c)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2274
                      "f contour_integrable_on linepath (shrink c) (shrink a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2275
          using fhp0  by (auto simp: valid_path_join dest: has_contour_integral_integrable)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2276
        have cmod_shr: "\<And>x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2277
          using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2278
        have sh_eq: "\<And>a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2279
          by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2280
        have "cmod y / (24 * C) \<le> cmod y / cmod (b - a) / 12"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2281
          using False \<open>C>0\<close> diff_2C [of b a] ynz
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2282
          by (auto simp: divide_simps hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2283
        have less_C: "\<lbrakk>u \<in> convex hull {a, b, c}; 0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> x * cmod u < C" for x u
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2284
          apply (cases "x=0", simp add: \<open>0<C\<close>)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2285
          using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2286
        { fix u v
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2287
          assume uv: "u \<in> convex hull {a, b, c}" "v \<in> convex hull {a, b, c}" "u\<noteq>v"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2288
             and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2289
          have shr_uv: "shrink u \<in> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2290
                       "shrink v \<in> interior(convex hull {a,b,c})"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2291
            using d e uv
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2292
            by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2293
          have cmod_fuv: "\<And>x. 0\<le>x \<Longrightarrow> x\<le>1 \<Longrightarrow> cmod (f (linepath (shrink u) (shrink v) x)) \<le> B"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2294
            using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2295
          have By_uv: "B * (12 * (e * cmod (u - v))) \<le> cmod y"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2296
            apply (rule order_trans [OF _ eCB])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2297
            using e \<open>B>0\<close> diff_2C [of u v] uv
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2298
            by (auto simp: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2299
          { fix x::real   assume x: "0\<le>x" "x\<le>1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2300
            have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2301
              apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2302
              using uv x d interior_subset
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2303
              apply (auto simp: hull_inc intro!: less_C)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2304
              done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2305
            have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2306
              by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2307
            have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2308
              using \<open>e>0\<close>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2309
              apply (simp add: ll norm_mult scaleR_diff_right)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2310
              apply (rule less_le_trans [OF _ e_le_d1])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2311
              using cmod_less_4C
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2312
              apply (force intro: norm_triangle_lt)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2313
              done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2314
            have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2315
              using x uv shr_uv cmod_less_dt
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2316
              by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2317
            also have "... \<le> cmod y / cmod (v - u) / 12"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2318
              using False uv \<open>C>0\<close> diff_2C [of v u] ynz
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2319
              by (auto simp: divide_simps hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2320
            finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \<le> cmod y / cmod (v - u) / 12"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2321
              by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2322
            then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \<le> cmod y"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2323
              using uv False by (auto simp: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2324
            have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2325
                  cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2326
                  \<le> cmod y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2327
              apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2328
              apply (rule add_mono [OF mult_mono])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2329
              using By_uv e \<open>0 < B\<close> \<open>0 < C\<close> x ynz
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2330
              apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2331
              apply (simp add: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2332
              done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2333
          } note cmod_diff_le = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2334
          have f_uv: "continuous_on (closed_segment u v) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2335
            by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2336
          have **: "\<And>f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2337
            by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2338
          have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \<le> norm y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2339
            apply (rule order_trans)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2340
            apply (rule has_integral_bound
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2341
                    [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2342
                        "\<lambda>x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2343
                        _ 0 1 ])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2344
            using ynz \<open>0 < B\<close> \<open>0 < C\<close>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2345
            apply (simp_all del: le_divide_eq_numeral1)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2346
            apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2347
                             fpi_uv f_uv contour_integrable_continuous_linepath, clarify)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2348
            apply (simp only: **)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2349
            apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2350
            done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2351
          } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2352
          have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \<le> norm y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2353
            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2354
          moreover
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2355
          have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \<le> norm y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2356
            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2357
          moreover
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2358
          have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \<le> norm y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2359
            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2360
          ultimately
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2361
          have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) +
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2362
                     (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2363
                \<le> norm y / 6 + norm y / 6 + norm y / 6"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2364
            by (metis norm_triangle_le add_mono)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2365
          also have "... = norm y / 2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2366
            by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2367
          finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2368
                          (?pathint a b + ?pathint b c + ?pathint c a))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2369
                \<le> norm y / 2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2370
            by (simp add: algebra_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2371
          then
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2372
          have "norm(?pathint a b + ?pathint b c + ?pathint c a) \<le> norm y / 2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2373
            by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2374
          then have "False"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2375
            using pi_eq_y ynz by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2376
        }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2377
        moreover have "uniformly_continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2378
          by (simp add: contf compact_convex_hull compact_uniformly_continuous)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2379
        ultimately have "False"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2380
          unfolding uniformly_continuous_on_def
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2381
          by (force simp: ynz \<open>0 < C\<close> dist_norm)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2382
        then show ?thesis ..
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2383
      qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2384
  }
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2385
  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2386
    using fabc contour_integrable_continuous_linepath by auto
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2387
  ultimately show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2388
    using has_contour_integral_integral by fastforce
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2389
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2390
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2391
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2392
subsection\<open>Version allowing finite number of exceptional points\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2393
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2394
lemma Cauchy_theorem_triangle_cofinite:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2395
  assumes "continuous_on (convex hull {a,b,c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2396
      and "finite s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2397
      and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - s \<Longrightarrow> f complex_differentiable (at x))"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2398
     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2399
using assms
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2400
proof (induction "card s" arbitrary: a b c s rule: less_induct)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2401
  case (less s a b c)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2402
  show ?case
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2403
  proof (cases "s={}")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2404
    case True with less show ?thesis
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  2405
      by (fastforce simp: holomorphic_on_def complex_differentiable_at_within
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2406
                    Cauchy_theorem_triangle_interior)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2407
  next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2408
    case False
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2409
    then obtain d s' where d: "s = insert d s'" "d \<notin> s'"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2410
      by (meson Set.set_insert all_not_in_conv)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2411
    then show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2412
    proof (cases "d \<in> convex hull {a,b,c}")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2413
      case False
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2414
      show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2415
        apply (rule less.hyps [of "s'"])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2416
        using False d \<open>finite s\<close> interior_subset
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2417
        apply (auto intro!: less.prems)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2418
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2419
    next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2420
      case True
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2421
      have *: "convex hull {a, b, d} \<subseteq> convex hull {a, b, c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2422
        by (meson True hull_subset insert_subset convex_hull_subset)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2423
      have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2424
        apply (rule less.hyps [of "s'"])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2425
        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2426
        apply (auto intro!: less.prems continuous_on_subset [OF  _ *])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2427
        apply (metis * insert_absorb insert_subset interior_mono)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2428
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2429
      have *: "convex hull {b, c, d} \<subseteq> convex hull {a, b, c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2430
        by (meson True hull_subset insert_subset convex_hull_subset)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2431
      have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2432
        apply (rule less.hyps [of "s'"])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2433
        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2434
        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2435
        apply (metis * insert_absorb insert_subset interior_mono)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2436
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2437
      have *: "convex hull {c, a, d} \<subseteq> convex hull {a, b, c}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2438
        by (meson True hull_subset insert_subset convex_hull_subset)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2439
      have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2440
        apply (rule less.hyps [of "s'"])
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2441
        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2442
        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2443
        apply (metis * insert_absorb insert_subset interior_mono)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2444
        done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2445
      have "f contour_integrable_on linepath a b"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2446
        using less.prems
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2447
        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2448
      moreover have "f contour_integrable_on linepath b c"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2449
        using less.prems
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2450
        by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3))
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2451
      moreover have "f contour_integrable_on linepath c a"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2452
        using less.prems
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2453
        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2454
      ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2455
        by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2456
      { fix y::complex
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2457
        assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2458
           and ynz: "y \<noteq> 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2459
        have cont_ad: "continuous_on (closed_segment a d) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2460
          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2461
        have cont_bd: "continuous_on (closed_segment b d) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2462
          by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2463
        have cont_cd: "continuous_on (closed_segment c d) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2464
          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2))
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2465
        have "contour_integral  (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2466
                "contour_integral  (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2467
                "contour_integral  (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2468
            using has_chain_integral_chain_integral3 [OF abd]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2469
                  has_chain_integral_chain_integral3 [OF bcd]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2470
                  has_chain_integral_chain_integral3 [OF cad]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2471
            by (simp_all add: algebra_simps add_eq_0_iff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2472
        then have ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2473
          using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2474
      }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2475
      then show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2476
        using fpi contour_integrable_on_def by blast
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2477
    qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2478
  qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2479
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2480
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2481
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2482
subsection\<open>Cauchy's theorem for an open starlike set\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2483
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2484
lemma starlike_convex_subset:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2485
  assumes s: "a \<in> s" "closed_segment b c \<subseteq> s" and subs: "\<And>x. x \<in> s \<Longrightarrow> closed_segment a x \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2486
    shows "convex hull {a,b,c} \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2487
      using s
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2488
      apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull)
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  2489
      apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2490
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2491
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2492
lemma triangle_contour_integrals_starlike_primitive:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2493
  assumes contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2494
      and s: "a \<in> s" "open s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2495
      and x: "x \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2496
      and subs: "\<And>y. y \<in> s \<Longrightarrow> closed_segment a y \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2497
      and zer: "\<And>b c. closed_segment b c \<subseteq> s
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2498
                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2499
                       contour_integral (linepath c a) f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2500
    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2501
proof -
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2502
  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2503
  { fix e y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2504
    assume e: "0 < e" and bxe: "ball x e \<subseteq> s" and close: "cmod (y - x) < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2505
    have y: "y \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2506
      using bxe close  by (force simp: dist_norm norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2507
    have cont_ayf: "continuous_on (closed_segment a y) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2508
      using contf continuous_on_subset subs y by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2509
    have xys: "closed_segment x y \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2510
      apply (rule order_trans [OF _ bxe])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2511
      using close
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2512
      by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2513
    have "?pathint a y - ?pathint a x = ?pathint x y"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2514
      using zer [OF xys]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2515
  } note [simp] = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2516
  { fix e::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2517
    assume e: "0 < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2518
    have cont_atx: "continuous (at x) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2519
      using x s contf continuous_on_eq_continuous_at by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2520
    then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. cmod (y - x) < d1 \<Longrightarrow> cmod (f y - f x) < e/2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2521
      unfolding continuous_at Lim_at dist_norm  using e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2522
      by (drule_tac x="e/2" in spec) force
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2523
    obtain d2 where d2: "d2>0" "ball x d2 \<subseteq> s" using  \<open>open s\<close> x
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2524
      by (auto simp: open_contains_ball)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2525
    have dpos: "min d1 d2 > 0" using d1 d2 by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2526
    { fix y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2527
      assume yx: "y \<noteq> x" and close: "cmod (y - x) < min d1 d2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2528
      have y: "y \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2529
        using d2 close  by (force simp: dist_norm norm_minus_commute)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2530
      have fxy: "f contour_integrable_on linepath x y"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2531
        apply (rule contour_integrable_continuous_linepath)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2532
        apply (rule continuous_on_subset [OF contf])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2533
        using close d2
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2534
        apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2535
        done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2536
      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2537
        by (auto simp: contour_integrable_on_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2538
      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2539
        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2540
      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2541
        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2542
        using e apply simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2543
        apply (rule d1_less [THEN less_imp_le])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2544
        using close segment_bound
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2545
        apply force
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2546
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2547
      also have "... < e * cmod (y - x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2548
        by (simp add: e yx)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2549
      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2550
        using i yx  by (simp add: contour_integral_unique divide_less_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2551
    }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2552
    then have "\<exists>d>0. \<forall>y. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2553
      using dpos by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2554
  }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2555
  then have *: "(\<lambda>y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) -- x --> 0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2556
    by (simp add: Lim_at dist_norm inverse_eq_divide)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2557
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2558
    apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2559
    apply (rule Lim_transform [OF * Lim_eventually])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2560
    apply (simp add: inverse_eq_divide [symmetric] eventually_at)
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2561
    using \<open>open s\<close> x
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2562
    apply (force simp: dist_norm open_contains_ball)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2563
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2564
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2565
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2566
(** Existence of a primitive.*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2567
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2568
lemma holomorphic_starlike_primitive:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2569
  assumes contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2570
      and s: "starlike s" and os: "open s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2571
      and k: "finite k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2572
      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f complex_differentiable at x"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2573
    shows "\<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2574
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2575
  obtain a where a: "a\<in>s" and a_cs: "\<And>x. x\<in>s \<Longrightarrow> closed_segment a x \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2576
    using s by (auto simp: starlike_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2577
  { fix x b c
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2578
    assume "x \<in> s" "closed_segment b c \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2579
    then have abcs: "convex hull {a, b, c} \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2580
      by (simp add: a a_cs starlike_convex_subset)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2581
    then have *: "continuous_on (convex hull {a, b, c}) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2582
      by (simp add: continuous_on_subset [OF contf])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2583
    have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2584
      apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2585
      using abcs apply (simp add: continuous_on_subset [OF contf])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2586
      using * abcs interior_subset apply (auto intro: fcd)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2587
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2588
  } note 0 = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2589
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2590
    apply (intro exI ballI)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2591
    apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2592
    apply (metis a_cs)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2593
    apply (metis has_chain_integral_chain_integral3 0)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2594
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2595
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2596
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2597
lemma Cauchy_theorem_starlike:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2598
 "\<lbrakk>open s; starlike s; finite k; continuous_on s f;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2599
   \<And>x. x \<in> s - k \<Longrightarrow> f complex_differentiable at x;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2600
   valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2601
   \<Longrightarrow> (f has_contour_integral 0)  g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2602
  by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2603
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2604
lemma Cauchy_theorem_starlike_simple:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2605
  "\<lbrakk>open s; starlike s; f holomorphic_on s; valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2606
   \<Longrightarrow> (f has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2607
apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2608
apply (simp_all add: holomorphic_on_imp_continuous_on)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2609
apply (metis at_within_open holomorphic_on_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2610
done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2611
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2612
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2613
subsection\<open>Cauchy's theorem for a convex set\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2614
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2615
text\<open>For a convex set we can avoid assuming openness and boundary analyticity\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2616
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2617
lemma triangle_contour_integrals_convex_primitive:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2618
  assumes contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2619
      and s: "a \<in> s" "convex s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2620
      and x: "x \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2621
      and zer: "\<And>b c. \<lbrakk>b \<in> s; c \<in> s\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2622
                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2623
                       contour_integral (linepath c a) f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2624
    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2625
proof -
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2626
  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2627
  { fix y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2628
    assume y: "y \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2629
    have cont_ayf: "continuous_on (closed_segment a y) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2630
      using s y  by (meson contf continuous_on_subset convex_contains_segment)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2631
    have xys: "closed_segment x y \<subseteq> s"  (*?*)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2632
      using convex_contains_segment s x y by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2633
    have "?pathint a y - ?pathint a x = ?pathint x y"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2634
      using zer [OF x y]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2635
  } note [simp] = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2636
  { fix e::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2637
    assume e: "0 < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2638
    have cont_atx: "continuous (at x within s) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2639
      using x s contf  by (simp add: continuous_on_eq_continuous_within)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2640
    then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. \<lbrakk>y \<in> s; cmod (y - x) < d1\<rbrakk> \<Longrightarrow> cmod (f y - f x) < e/2"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2641
      unfolding continuous_within Lim_within dist_norm using e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2642
      by (drule_tac x="e/2" in spec) force
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2643
    { fix y
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2644
      assume yx: "y \<noteq> x" and close: "cmod (y - x) < d1" and y: "y \<in> s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2645
      have fxy: "f contour_integrable_on linepath x y"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2646
        using convex_contains_segment s x y
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2647
        by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2648
      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2649
        by (auto simp: contour_integrable_on_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2650
      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2651
        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2652
      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2653
        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2654
        using e apply simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2655
        apply (rule d1_less [THEN less_imp_le])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2656
        using convex_contains_segment s(2) x y apply blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2657
        using close segment_bound(1) apply fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2658
        done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2659
      also have "... < e * cmod (y - x)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2660
        by (simp add: e yx)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2661
      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2662
        using i yx  by (simp add: contour_integral_unique divide_less_eq)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2663
    }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2664
    then have "\<exists>d>0. \<forall>y\<in>s. y \<noteq> x \<and> cmod (y-x) < d \<longrightarrow> cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2665
      using d1 by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2666
  }
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2667
  then have *: "((\<lambda>y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) ---> 0) (at x within s)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2668
    by (simp add: Lim_within dist_norm inverse_eq_divide)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2669
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2670
    apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2671
    apply (rule Lim_transform [OF * Lim_eventually])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2672
    using linordered_field_no_ub
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2673
    apply (force simp: inverse_eq_divide [symmetric] eventually_at)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2674
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2675
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2676
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2677
lemma pathintegral_convex_primitive:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2678
  "\<lbrakk>convex s; continuous_on s f;
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2679
    \<And>a b c. \<lbrakk>a \<in> s; b \<in> s; c \<in> s\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)\<rbrakk>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2680
         \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2681
  apply (cases "s={}")
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2682
  apply (simp_all add: ex_in_conv [symmetric])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2683
  apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2684
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2685
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2686
lemma holomorphic_convex_primitive:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2687
  "\<lbrakk>convex s; finite k; continuous_on s f;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2688
    \<And>x. x \<in> interior s - k \<Longrightarrow> f complex_differentiable at x\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2689
   \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2690
apply (rule pathintegral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2691
prefer 3
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2692
apply (erule continuous_on_subset)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2693
apply (simp add: subset_hull continuous_on_subset, assumption+)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2694
by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2695
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2696
lemma Cauchy_theorem_convex:
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2697
    "\<lbrakk>continuous_on s f; convex s; finite k;
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2698
      \<And>x. x \<in> interior s - k \<Longrightarrow> f complex_differentiable at x;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2699
     valid_path g; path_image g \<subseteq> s;
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2700
     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2701
  by (metis holomorphic_convex_primitive Cauchy_theorem_primitive)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2702
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2703
lemma Cauchy_theorem_convex_simple:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2704
    "\<lbrakk>f holomorphic_on s; convex s;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2705
     valid_path g; path_image g \<subseteq> s;
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2706
     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2707
  apply (rule Cauchy_theorem_convex)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2708
  apply (simp_all add: holomorphic_on_imp_continuous_on)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2709
  apply (rule finite.emptyI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2710
  using at_within_interior holomorphic_on_def interior_subset by fastforce
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2711
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2712
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2713
text\<open>In particular for a disc\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2714
lemma Cauchy_theorem_disc:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2715
    "\<lbrakk>finite k; continuous_on (cball a e) f;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2716
      \<And>x. x \<in> ball a e - k \<Longrightarrow> f complex_differentiable at x;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2717
     valid_path g; path_image g \<subseteq> cball a e;
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2718
     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2719
  apply (rule Cauchy_theorem_convex)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2720
  apply (auto simp: convex_cball interior_cball)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2721
  done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2722
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2723
lemma Cauchy_theorem_disc_simple:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2724
    "\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e;
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2725
     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2726
by (simp add: Cauchy_theorem_convex_simple)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2727
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2728
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2729
subsection\<open>Generalize integrability to local primitives\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2730
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2731
lemma contour_integral_local_primitive_lemma:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2732
  fixes f :: "complex\<Rightarrow>complex"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2733
  shows
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2734
    "\<lbrakk>g piecewise_differentiable_on {a..b};
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2735
      \<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s);
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2736
      \<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s\<rbrakk>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2737
     \<Longrightarrow> (\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b}))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2738
            integrable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2739
  apply (cases "cbox a b = {}", force)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2740
  apply (simp add: integrable_on_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2741
  apply (rule exI)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2742
  apply (rule contour_integral_primitive_lemma, assumption+)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2743
  using atLeastAtMost_iff by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2744
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2745
lemma contour_integral_local_primitive_any:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2746
  fixes f :: "complex \<Rightarrow> complex"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2747
  assumes gpd: "g piecewise_differentiable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2748
      and dh: "\<And>x. x \<in> s
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2749
               \<Longrightarrow> \<exists>d h. 0 < d \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2750
                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2751
      and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2752
  shows "(\<lambda>x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2753
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2754
  { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2755
    assume x: "a \<le> x" "x \<le> b"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2756
    obtain d h where d: "0 < d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2757
               and h: "(\<And>y. norm(y - g x) < d \<Longrightarrow> (h has_field_derivative f y) (at y within s))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2758
      using x gs dh by (metis atLeastAtMost_iff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2759
    have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2760
    then obtain e where e: "e>0" and lessd: "\<And>x'. x' \<in> {a..b} \<Longrightarrow> \<bar>x' - x\<bar> < e \<Longrightarrow> cmod (g x' - g x) < d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2761
      using x d
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2762
      apply (auto simp: dist_norm continuous_on_iff)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2763
      apply (drule_tac x=x in bspec)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2764
      using x apply simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2765
      apply (drule_tac x=d in spec, auto)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2766
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2767
    have "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> {u..v} \<subseteq> ball x d \<and> (u \<le> v \<longrightarrow> a \<le> u \<and> v \<le> b) \<longrightarrow>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2768
                          (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2769
      apply (rule_tac x=e in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2770
      using e
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2771
      apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2772
      apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2773
        apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2774
       apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2775
      done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2776
  } then
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2777
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2778
    by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2779
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2780
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2781
lemma contour_integral_local_primitive:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2782
  fixes f :: "complex \<Rightarrow> complex"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2783
  assumes g: "valid_path g" "path_image g \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2784
      and dh: "\<And>x. x \<in> s
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2785
               \<Longrightarrow> \<exists>d h. 0 < d \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2786
                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2787
  shows "f contour_integrable_on g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2788
  using g
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2789
  apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2790
            has_integral_localized_vector_derivative integrable_on_def [symmetric])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2791
  using contour_integral_local_primitive_any [OF _ dh]
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  2792
  by (meson image_subset_iff piecewise_C1_imp_differentiable)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2793
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2794
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2795
text\<open>In particular if a function is holomorphic\<close>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2796
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2797
lemma contour_integrable_holomorphic:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2798
  assumes contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2799
      and os: "open s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2800
      and k: "finite k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2801
      and g: "valid_path g" "path_image g \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2802
      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f complex_differentiable at x"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2803
    shows "f contour_integrable_on g"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2804
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2805
  { fix z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2806
    assume z: "z \<in> s"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2807
    obtain d where d: "d>0" "ball z d \<subseteq> s" using  \<open>open s\<close> z
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2808
      by (auto simp: open_contains_ball)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2809
    then have contfb: "continuous_on (ball z d) f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2810
      using contf continuous_on_subset by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2811
    obtain h where "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within ball z d)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2812
      using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2813
            interior_subset by force
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2814
    then have "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within s)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2815
      by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2816
    then have "\<exists>h. (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2817
      by (force simp: dist_norm norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2818
    then have "\<exists>d h. 0 < d \<and> (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2819
      using d by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2820
  }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2821
  then show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2822
    by (rule contour_integral_local_primitive [OF g])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2823
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2824
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2825
lemma contour_integrable_holomorphic_simple:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2826
  assumes contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2827
      and os: "open s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2828
      and g: "valid_path g" "path_image g \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2829
      and fh: "f holomorphic_on s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2830
    shows "f contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2831
  apply (rule contour_integrable_holomorphic [OF contf os Finite_Set.finite.emptyI g])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2832
  using fh  by (simp add: complex_differentiable_def holomorphic_on_open os)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2833
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  2834
lemma continuous_on_inversediff:
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  2835
  fixes z:: "'a::real_normed_field" shows "z \<notin> s \<Longrightarrow> continuous_on s (\<lambda>w. 1 / (w - z))"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  2836
  by (rule continuous_intros | force)+
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  2837
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2838
corollary contour_integrable_inversediff:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2839
    "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) contour_integrable_on g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2840
apply (rule contour_integrable_holomorphic_simple [of "UNIV-{z}", OF continuous_on_inversediff])
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  2841
apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2842
done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2843
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2844
text\<open>Key fact that path integral is the same for a "nearby" path. This is the
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2845
 main lemma for the homotopy form of Cauchy's theorem and is also useful
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2846
 if we want "without loss of generality" to assume some nice properties of a
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2847
 path (e.g. smoothness). It can also be used to define the integrals of
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2848
 analytic functions over arbitrary continuous paths. This is just done for
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2849
 winding numbers now.
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2850
\<close>
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2851
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2852
text\<open>A technical definition to avoid duplication of similar proofs,
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2853
     for paths joined at the ends versus looping paths\<close>
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2854
definition linked_paths :: "bool \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2855
  where "linked_paths atends g h ==
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2856
        (if atends then pathstart h = pathstart g \<and> pathfinish h = pathfinish g
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2857
                   else pathfinish g = pathstart g \<and> pathfinish h = pathstart h)"
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2858
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2859
text\<open>This formulation covers two cases: @{term g} and @{term h} share their
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2860
      start and end points; @{term g} and @{term h} both loop upon themselves.\<close>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2861
lemma contour_integral_nearby:
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2862
  assumes os: "open s" and p: "path p" "path_image p \<subseteq> s"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2863
    shows
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2864
       "\<exists>d. 0 < d \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2865
            (\<forall>g h. valid_path g \<and> valid_path h \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2866
                  (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2867
                  linked_paths atends g h
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2868
                  \<longrightarrow> path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2869
                      (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2870
proof -
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2871
  have "\<forall>z. \<exists>e. z \<in> path_image p \<longrightarrow> 0 < e \<and> ball z e \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2872
    using open_contains_ball os p(2) by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2873
  then obtain ee where ee: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2874
    by metis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2875
  def cover \<equiv> "(\<lambda>z. ball z (ee z/3)) ` (path_image p)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2876
  have "compact (path_image p)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2877
    by (metis p(1) compact_path_image)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2878
  moreover have "path_image p \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2879
    using ee by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2880
  ultimately have "\<exists>D \<subseteq> cover. finite D \<and> path_image p \<subseteq> \<Union>D"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2881
    by (simp add: compact_eq_heine_borel cover_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2882
  then obtain D where D: "D \<subseteq> cover" "finite D" "path_image p \<subseteq> \<Union>D"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2883
    by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2884
  then obtain k where k: "k \<subseteq> {0..1}" "finite k" and D_eq: "D = ((\<lambda>z. ball z (ee z / 3)) \<circ> p) ` k"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2885
    apply (simp add: cover_def path_image_def image_comp)
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2886
    apply (blast dest!: finite_subset_image [OF \<open>finite D\<close>])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2887
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2888
  then have kne: "k \<noteq> {}"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2889
    using D by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2890
  have pi: "\<And>i. i \<in> k \<Longrightarrow> p i \<in> path_image p"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2891
    using k  by (auto simp: path_image_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2892
  then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2893
    by (metis ee)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2894
  def e \<equiv> "Min((ee o p) ` k)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2895
  have fin_eep: "finite ((ee o p) ` k)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2896
    using k  by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2897
  have enz: "0 < e"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2898
    using ee k  by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2899
  have "uniformly_continuous_on {0..1} p"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2900
    using p  by (simp add: path_def compact_uniformly_continuous)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2901
  then obtain d::real where d: "d>0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2902
          and de: "\<And>x x'. \<bar>x' - x\<bar> < d \<Longrightarrow> x\<in>{0..1} \<Longrightarrow> x'\<in>{0..1} \<Longrightarrow> cmod (p x' - p x) < e/3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2903
    unfolding uniformly_continuous_on_def dist_norm real_norm_def
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2904
    by (metis divide_pos_pos enz zero_less_numeral)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2905
  then obtain N::nat where N: "N>0" "inverse N < d"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2906
    using real_arch_inv [of d]   by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2907
  { fix g h
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2908
    assume g: "valid_path g" and gp: "\<forall>t\<in>{0..1}. cmod (g t - p t) < e / 3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2909
       and h: "valid_path h" and hp: "\<forall>t\<in>{0..1}. cmod (h t - p t) < e / 3"
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  2910
       and joins: "linked_paths atends g h"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2911
    { fix t::real
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2912
      assume t: "0 \<le> t" "t \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2913
      then obtain u where u: "u \<in> k" and ptu: "p t \<in> ball(p u) (ee(p u) / 3)"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2914
        using \<open>path_image p \<subseteq> \<Union>D\<close> D_eq by (force simp: path_image_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2915
      then have ele: "e \<le> ee (p u)" using fin_eep
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2916
        by (simp add: e_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2917
      have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2918
        using gp hp t by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2919
      with ele have "cmod (g t - p t) < ee (p u) / 3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2920
                    "cmod (h t - p t) < ee (p u) / 3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2921
        by linarith+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2922
      then have "g t \<in> ball(p u) (ee(p u))"  "h t \<in> ball(p u) (ee(p u))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2923
        using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2924
              norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  2925
        by (force simp: dist_norm ball_def norm_minus_commute)+
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2926
      then have "g t \<in> s" "h t \<in> s" using ee u k
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2927
        by (auto simp: path_image_def ball_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2928
    }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2929
    then have ghs: "path_image g \<subseteq> s" "path_image h \<subseteq> s"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2930
      by (auto simp: path_image_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2931
    moreover
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2932
    { fix f
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2933
      assume fhols: "f holomorphic_on s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2934
      then have fpa: "f contour_integrable_on g"  "f contour_integrable_on h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2935
        using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2936
        by blast+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2937
      have contf: "continuous_on s f"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2938
        by (simp add: fhols holomorphic_on_imp_continuous_on)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2939
      { fix z
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2940
        assume z: "z \<in> path_image p"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2941
        have "f holomorphic_on ball z (ee z)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2942
          using fhols ee z holomorphic_on_subset by blast
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2943
        then have "\<exists>ff. (\<forall>w \<in> ball z (ee z). (ff has_field_derivative f w) (at w))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2944
          using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified]
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2945
          by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2946
      }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2947
      then obtain ff where ff:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2948
            "\<And>z w. \<lbrakk>z \<in> path_image p; w \<in> ball z (ee z)\<rbrakk> \<Longrightarrow> (ff z has_field_derivative f w) (at w)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2949
        by metis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2950
      { fix n
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2951
        assume n: "n \<le> N"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2952
        then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  2953
                   contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2954
        proof (induct n)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2955
          case 0 show ?case by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2956
        next
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2957
          case (Suc n)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2958
          obtain t where t: "t \<in> k" and "p (n/N) \<in> ball(p t) (ee(p t) / 3)"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2959
            using \<open>path_image p \<subseteq> \<Union>D\<close> [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  2960
            by (force simp: path_image_def)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2961
          then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2962
            by (simp add: dist_norm)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2963
          have e3le: "e/3 \<le> ee (p t) / 3"  using fin_eep t
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2964
            by (simp add: e_def)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2965
          { fix x
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2966
            assume x: "n/N \<le> x" "x \<le> (1 + n)/N"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2967
            then have nN01: "0 \<le> n/N" "(1 + n)/N \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2968
              using Suc.prems by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2969
            then have x01: "0 \<le> x" "x \<le> 1"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2970
              using x by linarith+
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2971
            have "cmod (p t - p x)  < ee (p t) / 3 + e/3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2972
              apply (rule norm_diff_triangle_less [OF ptu de])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2973
              using x N x01 Suc.prems
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2974
              apply (auto simp: field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2975
              done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2976
            then have ptx: "cmod (p t - p x) < 2*ee (p t)/3"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2977
              using e3le eepi [OF t] by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2978
            have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 "
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2979
              apply (rule norm_diff_triangle_less [OF ptx])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2980
              using gp x01 by (simp add: norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2981
            also have "... \<le> ee (p t)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2982
              using e3le eepi [OF t] by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2983
            finally have gg: "cmod (p t - g x) < ee (p t)" .
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2984
            have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 "
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2985
              apply (rule norm_diff_triangle_less [OF ptx])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2986
              using hp x01 by (simp add: norm_minus_commute)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2987
            also have "... \<le> ee (p t)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2988
              using e3le eepi [OF t] by simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2989
            finally have "cmod (p t - g x) < ee (p t)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2990
                         "cmod (p t - h x) < ee (p t)"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2991
              using gg by auto
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2992
          } note ptgh_ee = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2993
          have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \<subseteq> ball (p t) (ee (p t))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2994
            using ptgh_ee [of "n/N"] Suc.prems
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: 61520
diff changeset
  2995
            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2996
          then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \<subseteq> s"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  2997
            using \<open>N>0\<close> Suc.prems
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: 61520
diff changeset
  2998
            apply (simp add: path_image_join field_simps closed_segment_commute)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2999
            apply (erule order_trans)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3000
            apply (simp add: ee pi t)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3001
            done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3002
          have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N)))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3003
                  \<subseteq> ball (p t) (ee (p t))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3004
            using ptgh_ee [of "(1+n)/N"] Suc.prems
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: 61520
diff changeset
  3005
            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3006
          then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \<subseteq> s"
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  3007
            using \<open>N>0\<close> Suc.prems ee pi t
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3008
            by (auto simp: Path_Connected.path_image_join field_simps)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3009
          have pi_subset_ball:
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3010
                "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3011
                             subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N)))
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3012
                 \<subseteq> ball (p t) (ee (p t))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3013
            apply (intro subset_path_image_join pi_hgn pi_ghn')
61222
05d28dc76e5c isabelle update_cartouches;
wenzelm
parents: 61204
diff changeset
  3014
            using \<open>N>0\<close> Suc.prems
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3015
            apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3016
            done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3017
          have pi0: "(f has_contour_integral 0)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3018
                       (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3019
                        subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3020
            apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3021
            apply (metis ff open_ball at_within_open pi t)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3022
            apply (intro valid_path_join)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3023
            using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3024
            done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3025
          have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3026
            using Suc.prems by (simp add: contour_integrable_subpath g fpa)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3027
          have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3028
            using gh_n's
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3029
            by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3030
          have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3031
            using gh_ns
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3032
            by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3033
          have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3034
                     contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3035
                     contour_integral (subpath ((Suc n) / N) (n/N) h) f +
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3036
                     contour_integral (linepath (h (n/N)) (g (n/N))) f = 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3037
            using contour_integral_unique [OF pi0] Suc.prems
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3038
            by (simp add: g h fpa valid_path_subpath contour_integrable_subpath
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: 61520
diff changeset
  3039
                          fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3040
          have *: "\<And>hn he hn' gn gd gn' hgn ghn gh0 ghn'.
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3041
                    \<lbrakk>hn - gn = ghn - gh0;
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3042
                     gd + ghn' + he + hgn = (0::complex);
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3043
                     hn - he = hn'; gn + gd = gn'; hgn = -ghn\<rbrakk> \<Longrightarrow> hn' - gn' = ghn' - gh0"
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3044
            by (auto simp: algebra_simps)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3045
          have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3046
                contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3047
            unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"]
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3048
            using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3049
          also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3050
            using Suc.prems by (simp add: contour_integral_subpath_combine h fpa)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3051
          finally have pi0_eq:
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3052
               "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3053
                contour_integral (subpath 0 ((Suc n) / N) h) f" .
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3054
          show ?case
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3055
            apply (rule * [OF Suc.hyps eq0 pi0_eq])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3056
            using Suc.prems
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3057
            apply (simp_all add: g h fpa contour_integral_subpath_combine
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3058
                     contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3059
                     continuous_on_subset [OF contf gh_ns])
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3060
            done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3061
      qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3062
      } note ind = this
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3063
      have "contour_integral h f = contour_integral g f"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3064
        using ind [OF order_refl] N joins
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  3065
        by (simp add: linked_paths_def pathstart_def pathfinish_def split: split_if_asm)
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3066
    }
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3067
    ultimately
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3068
    have "path_image g \<subseteq> s \<and> path_image h \<subseteq> s \<and> (\<forall>f. f holomorphic_on s \<longrightarrow> contour_integral h f = contour_integral g f)"
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3069
      by metis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3070
  } note * = this
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3071
  show ?thesis
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3072
    apply (rule_tac x="e/3" in exI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3073
    apply (rule conjI)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3074
    using enz apply simp
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3075
    apply (clarsimp simp only: ball_conj_distrib)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3076
    apply (rule *; assumption)
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3077
    done
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3078
qed
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3079
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3080
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3081
lemma
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3082
  assumes "open s" "path p" "path_image p \<subseteq> s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3083
    shows contour_integral_nearby_ends:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3084
      "\<exists>d. 0 < d \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3085
              (\<forall>g h. valid_path g \<and> valid_path h \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3086
                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3087
                    pathstart h = pathstart g \<and> pathfinish h = pathfinish g
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3088
                    \<longrightarrow> path_image g \<subseteq> s \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3089
                        path_image h \<subseteq> s \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3090
                        (\<forall>f. f holomorphic_on s
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3091
                            \<longrightarrow> contour_integral h f = contour_integral g f))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3092
    and contour_integral_nearby_loops:
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3093
      "\<exists>d. 0 < d \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3094
              (\<forall>g h. valid_path g \<and> valid_path h \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3095
                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3096
                    pathfinish g = pathstart g \<and> pathfinish h = pathstart h
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3097
                    \<longrightarrow> path_image g \<subseteq> s \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3098
                        path_image h \<subseteq> s \<and>
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3099
                        (\<forall>f. f holomorphic_on s
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3100
                            \<longrightarrow> contour_integral h f = contour_integral g f))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3101
  using contour_integral_nearby [OF assms, where atends=True]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3102
  using contour_integral_nearby [OF assms, where atends=False]
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
  3103
  unfolding linked_paths_def by simp_all
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3104
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3105
corollary differentiable_polynomial_function:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3106
  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3107
  shows "polynomial_function p \<Longrightarrow> p differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3108
by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3109
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3110
lemma C1_differentiable_polynomial_function:
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3111
  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3112
  shows "polynomial_function p \<Longrightarrow> p C1_differentiable_on s"
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3113
  by (metis continuous_on_polymonial_function C1_differentiable_on_def  has_vector_derivative_polynomial_function)
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3114
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3115
lemma valid_path_polynomial_function:
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3116
  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3117
  shows "polynomial_function p \<Longrightarrow> valid_path p"
61190
2bd401e364f9 Massive revisions, as a valid path must now be continously differentiable (C!)
paulson <lp15@cam.ac.uk>
parents: 61104
diff changeset
  3118
by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3119
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3120
lemma valid_path_subpath_trivial [simp]:
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3121
    fixes g :: "real \<Rightarrow> 'a::euclidean_space"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3122
    shows "z \<noteq> g x \<Longrightarrow> valid_path (subpath x x g)"
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3123
  by (simp add: subpath_def valid_path_polynomial_function)
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
  3124
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3125
lemma contour_integral_bound_exists:
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3126
assumes s: "open s"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3127
    and g: "valid_path g"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3128
    and pag: "path_image g \<subseteq> s"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3129
  shows "\<exists>L. 0 < L \<and>
61200
a5674da43c2b eliminated hard tabs;
wenzelm
parents: 61190
diff changeset
  3130
       (\<forall>f B. f holomorphic_on s \<and> (\<forall>z \<in> s. norm(f z) \<le> B)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3131
         \<longrightarrow> norm(contour_integral g f) \<le> L*B)"
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3132
proof -
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3133
have "path g" using g
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3134
  by (simp add: valid_path_imp_path)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3135
then obtain d::real and p
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3136
  where d: "0 < d"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3137
    and p: "polynomial_function p" "path_image p \<subseteq> s"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3138
    and pi: "\<And>f. f holomorphic_on s \<Longrightarrow> contour_integral g f = contour_integral p f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3139
  using contour_integral_nearby_ends [OF s \<open>path g\<close> pag]
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3140
  apply clarify
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3141
  apply (drule_tac x=g in spec)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3142
  apply (simp only: assms)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3143
  apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3144
  done
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3145
then obtain p' where p': "polynomial_function p'"
61200
a5674da43c2b eliminated hard tabs;
wenzelm
parents: 61190
diff changeset
  3146
         "\<And>x. (p has_vector_derivative (p' x)) (at x)"
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3147
  using has_vector_derivative_polynomial_function by force
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3148
then have "bounded(p' ` {0..1})"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3149
  using continuous_on_polymonial_function
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3150
  by (force simp: intro!: compact_imp_bounded compact_continuous_image)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3151
then obtain L where L: "L>0" and nop': "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (p' x) \<le> L"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3152
  by (force simp: bounded_pos)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3153
{ fix f B
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3154
  assume f: "f holomorphic_on s"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3155
     and B: "\<And>z. z\<in>s \<Longrightarrow> cmod (f z) \<le> B"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3156
  then have "f contour_integrable_on p \<and> valid_path p"
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3157
    using p s
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3158
    by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on)
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3159
  moreover have "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (vector_derivative p (at x)) * cmod (f (p x)) \<le> L * B"
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3160
    apply (rule mult_mono)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3161
    apply (subst Derivative.vector_derivative_at; force intro: p' nop')
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3162
    using L B p
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3163
    apply (auto simp: path_image_def image_subset_iff)
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3164
    done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3165
  ultimately have "cmod (contour_integral g f) \<le> L * B"
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3166
    apply (simp add: pi [OF f])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3167
    apply (simp add: contour_integral_integral)
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3168
    apply (rule order_trans [OF integral_norm_bound_integral])
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3169
    apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult)
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3170
    done
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3171
} then
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3172
show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3173
  by (force simp: L contour_integral_integral)
61104
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3174
qed
3c2d4636cebc new lemmas about vector_derivative, complex numbers, paths, etc.
paulson
parents: 60809
diff changeset
  3175
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3176
subsection\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3177
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3178
text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3179
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3180
lemma continuous_disconnected_range_constant:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3181
  assumes s: "connected s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3182
      and conf: "continuous_on s f"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3183
      and fim: "f ` s \<subseteq> t"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3184
      and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3185
    shows "\<exists>a. \<forall>x \<in> s. f x = a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3186
proof (cases "s = {}")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3187
  case True then show ?thesis by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3188
next
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3189
  case False
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3190
  { fix x assume "x \<in> s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3191
    then have "f ` s \<subseteq> {f x}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3192
    by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3193
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3194
  with False show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3195
    by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3196
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3197
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3198
lemma discrete_subset_disconnected:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3199
  fixes s :: "'a::topological_space set"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3200
  fixes t :: "'b::real_normed_vector set"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3201
  assumes conf: "continuous_on s f"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3202
      and no: "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3203
   shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3204
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3205
  { fix x assume x: "x \<in> s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3206
    then obtain e where "e>0" and ele: "\<And>y. \<lbrakk>y \<in> s; f y \<noteq> f x\<rbrakk> \<Longrightarrow> e \<le> norm (f y - f x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3207
      using conf no [OF x] by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3208
    then have e2: "0 \<le> e / 2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3209
      by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3210
    have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3211
      apply (rule ccontr)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3212
      using connected_closed [of "connected_component_set (f ` s) (f x)"] `e>0`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3213
      apply (simp add: del: ex_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3214
      apply (drule spec [where x="cball (f x) (e / 2)"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3215
      apply (drule spec [where x="- ball(f x) e"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3216
      apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3217
        apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3218
       using centre_in_cball connected_component_refl_eq e2 x apply blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3219
      using ccs
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3220
      apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF `y \<in> s`])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3221
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3222
    moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3223
      by (auto simp: connected_component_in)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3224
    ultimately have "connected_component_set (f ` s) (f x) = {f x}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3225
      by (auto simp: x)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3226
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3227
  with assms show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3228
    by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3229
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3230
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3231
lemma finite_implies_discrete:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3232
  fixes s :: "'a::topological_space set"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3233
  assumes "finite (f ` s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3234
  shows "(\<forall>x \<in> s. \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3235
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3236
  have "\<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)" if "x \<in> s" for x
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3237
  proof (cases "f ` s - {f x} = {}")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3238
    case True
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3239
    with zero_less_numeral show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3240
      by (fastforce simp add: Set.image_subset_iff cong: conj_cong)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3241
  next
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3242
    case False
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3243
    then obtain z where z: "z \<in> s" "f z \<noteq> f x"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3244
      by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3245
    have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3246
      using assms by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3247
    then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3248
      apply (rule finite_imp_less_Inf)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3249
      using z apply force+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3250
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3251
    show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3252
      by (force intro!: * cInf_le_finite [OF finn])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3253
  qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3254
  with assms show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3255
    by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3256
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3257
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3258
text\<open>This proof requires the existence of two separate values of the range type.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3259
lemma finite_range_constant_imp_connected:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3260
  assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3261
              \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3262
    shows "connected s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3263
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3264
  { fix t u
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3265
    assume clt: "closedin (subtopology euclidean s) t"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3266
       and clu: "closedin (subtopology euclidean s) u"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3267
       and tue: "t \<inter> u = {}" and tus: "t \<union> u = s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3268
    have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3269
      apply (subst tus [symmetric])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3270
      apply (rule continuous_on_cases_local)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3271
      using clt clu tue
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3272
      apply (auto simp: tus continuous_on_const)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3273
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3274
    have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3275
      by (rule finite_subset [of _ "{0,1}"]) auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3276
    have "t = {} \<or> u = {}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3277
      using assms [OF conif fi] tus [symmetric]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3278
      by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3279
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3280
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3281
    by (simp add: connected_closed_in_eq)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3282
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3283
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3284
lemma continuous_disconnected_range_constant_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3285
      "(connected s \<longleftrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3286
           (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3287
            \<forall>t. continuous_on s f \<and> f ` s \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3288
            \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3289
  and continuous_discrete_range_constant_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3290
      "(connected s \<longleftrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3291
         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3292
          continuous_on s f \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3293
          (\<forall>x \<in> s. \<exists>e. 0 < e \<and> (\<forall>y. y \<in> s \<and> (f y \<noteq> f x) \<longrightarrow> e \<le> norm(f y - f x)))
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3294
          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis2)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3295
  and continuous_finite_range_constant_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3296
      "(connected s \<longleftrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3297
         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3298
          continuous_on s f \<and> finite (f ` s)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3299
          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis3)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3300
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3301
  have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3302
    \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3303
    by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3304
  have "?thesis1 \<and> ?thesis2 \<and> ?thesis3"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3305
    apply (rule *)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3306
    using continuous_disconnected_range_constant apply metis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3307
    apply clarify
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3308
    apply (frule discrete_subset_disconnected; blast)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3309
    apply (blast dest: finite_implies_discrete)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3310
    apply (blast intro!: finite_range_constant_imp_connected)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3311
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3312
  then show ?thesis1 ?thesis2 ?thesis3
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3313
    by blast+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3314
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3315
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3316
lemma continuous_discrete_range_constant:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3317
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3318
  assumes s: "connected s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3319
      and "continuous_on s f"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3320
      and "\<And>x. x \<in> s \<Longrightarrow> \<exists>e>0. \<forall>y. y \<in> s \<and> f y \<noteq> f x \<longrightarrow> e \<le> norm (f y - f x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3321
    shows "\<exists>a. \<forall>x \<in> s. f x = a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3322
  using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3323
  by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3324
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3325
lemma continuous_finite_range_constant:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3326
  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3327
  assumes "connected s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3328
      and "continuous_on s f"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3329
      and "finite (f ` s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3330
    shows "\<exists>a. \<forall>x \<in> s. f x = a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3331
  using assms continuous_finite_range_constant_eq
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3332
  by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3333
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3334
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3335
text\<open>We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3336
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3337
subsection\<open>Winding Numbers\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3338
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3339
text\<open>The result is an integer, but it doesn't have type @{typ int}!\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3340
definition winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3341
  "winding_number \<gamma> z \<equiv>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3342
    @n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3343
                    pathstart p = pathstart \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3344
                    pathfinish p = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3345
                    (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3346
                    contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3347
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3348
lemma winding_number:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3349
  assumes "path \<gamma>" "z \<notin> path_image \<gamma>" "0 < e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3350
    shows "\<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3351
               pathstart p = pathstart \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3352
               pathfinish p = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3353
               (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3354
               contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * winding_number \<gamma> z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3355
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3356
  have "path_image \<gamma> \<subseteq> UNIV - {z}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3357
    using assms by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3358
  then obtain d
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3359
    where d: "d>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3360
      and pi_eq: "\<And>h1 h2. valid_path h1 \<and> valid_path h2 \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3361
                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d) \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3362
                    pathstart h2 = pathstart h1 \<and> pathfinish h2 = pathfinish h1 \<longrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3363
                      path_image h1 \<subseteq> UNIV - {z} \<and> path_image h2 \<subseteq> UNIV - {z} \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3364
                      (\<forall>f. f holomorphic_on UNIV - {z} \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3365
    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3366
  then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3367
                          (\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3368
    using path_approx_polynomial_function [OF `path \<gamma>`, of "d/2"] d by auto
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3369
  def nn \<equiv> "1/(2* pi*ii) * contour_integral h (\<lambda>w. 1/(w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3370
  have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3371
                        pathstart p = pathstart \<gamma> \<and>  pathfinish p = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3372
                        (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3373
                        contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3374
                    (is "\<exists>n. \<forall>e > 0. ?PP e n")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3375
    proof (rule_tac x=nn in exI, clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3376
      fix e::real
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3377
      assume e: "e>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3378
      obtain p where p: "polynomial_function p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3379
            pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3380
        using path_approx_polynomial_function [OF `path \<gamma>`, of "min e (d/2)"] d `0<e` by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3381
      have "(\<lambda>w. 1 / (w - z)) holomorphic_on UNIV - {z}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3382
        by (auto simp: intro!: holomorphic_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3383
      then show "?PP e nn"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3384
        apply (rule_tac x=p in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3385
        using pi_eq [of h p] h p d
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3386
        apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3387
        done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3388
    qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3389
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3390
    unfolding winding_number_def
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3391
    apply (rule someI2_ex)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3392
    apply (blast intro: `0<e`)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3393
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3394
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3395
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3396
lemma winding_number_unique:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3397
  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3398
      and pi:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3399
        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3400
                          pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3401
                          (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3402
                          contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3403
   shows "winding_number \<gamma> z = n"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3404
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3405
  have "path_image \<gamma> \<subseteq> UNIV - {z}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3406
    using assms by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3407
  then obtain e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3408
    where e: "e>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3409
      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3410
                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3411
                    pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3412
                    contour_integral h2 f = contour_integral h1 f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3413
    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3414
  obtain p where p:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3415
     "valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3416
      pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3417
      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3418
      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3419
    using pi [OF e] by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3420
  obtain q where q:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3421
     "valid_path q \<and> z \<notin> path_image q \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3422
      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3423
      (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3424
    using winding_number [OF \<gamma> e] by blast
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3425
  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3426
    using p by auto
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3427
  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3428
    apply (rule pi_eq)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3429
    using p q
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3430
    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3431
  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3432
    using q by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3433
  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3434
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3435
    by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3436
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3437
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3438
lemma winding_number_unique_loop:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3439
  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3440
      and loop: "pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3441
      and pi:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3442
        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3443
                           pathfinish p = pathstart p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3444
                           (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3445
                           contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3446
   shows "winding_number \<gamma> z = n"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3447
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3448
  have "path_image \<gamma> \<subseteq> UNIV - {z}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3449
    using assms by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3450
  then obtain e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3451
    where e: "e>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3452
      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3453
                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3454
                    pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3455
                    contour_integral h2 f = contour_integral h1 f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3456
    using contour_integral_nearby_loops [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3457
  obtain p where p:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3458
     "valid_path p \<and> z \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3459
      pathfinish p = pathstart p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3460
      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3461
      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * ii * n"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3462
    using pi [OF e] by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3463
  obtain q where q:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3464
     "valid_path q \<and> z \<notin> path_image q \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3465
      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3466
      (\<forall>t\<in>{0..1}. cmod (\<gamma> t - q t) < e) \<and> contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3467
    using winding_number [OF \<gamma> e] by blast
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3468
  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3469
    using p by auto
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3470
  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3471
    apply (rule pi_eq)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3472
    using p q loop
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3473
    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3474
  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3475
    using q by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3476
  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3477
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3478
    by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3479
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3480
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3481
lemma winding_number_valid_path:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3482
  assumes "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3483
    shows "winding_number \<gamma> z = 1/(2*pi*ii) * contour_integral \<gamma> (\<lambda>w. 1/(w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3484
  using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3485
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3486
lemma has_contour_integral_winding_number:
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3487
  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3488
    shows "((\<lambda>w. 1/(w - z)) has_contour_integral (2*pi*ii*winding_number \<gamma> z)) \<gamma>"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3489
by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3490
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3491
lemma winding_number_trivial [simp]: "z \<noteq> a \<Longrightarrow> winding_number(linepath a a) z = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3492
  by (simp add: winding_number_valid_path)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3493
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3494
lemma winding_number_subpath_trivial [simp]: "z \<noteq> g x \<Longrightarrow> winding_number (subpath x x g) z = 0"
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3495
  by (simp add: path_image_subpath winding_number_valid_path)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3496
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3497
lemma winding_number_join:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3498
  assumes g1: "path g1" "z \<notin> path_image g1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3499
      and g2: "path g2" "z \<notin> path_image g2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3500
      and "pathfinish g1 = pathstart g2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3501
    shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3502
  apply (rule winding_number_unique)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3503
  using assms apply (simp_all add: not_in_path_image_join)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3504
  apply (frule winding_number [OF g2])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3505
  apply (frule winding_number [OF g1], clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3506
  apply (rename_tac p2 p1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3507
  apply (rule_tac x="p1+++p2" in exI)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3508
  apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3509
  apply (auto simp: joinpaths_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3510
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3511
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3512
lemma winding_number_reversepath:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3513
  assumes "path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3514
    shows "winding_number(reversepath \<gamma>) z = - (winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3515
  apply (rule winding_number_unique)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3516
  using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3517
  apply simp_all
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3518
  apply (frule winding_number [OF assms], clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3519
  apply (rule_tac x="reversepath p" in exI)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3520
  apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3521
  apply (auto simp: reversepath_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3522
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3523
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3524
lemma winding_number_shiftpath:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3525
  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3526
      and "pathfinish \<gamma> = pathstart \<gamma>" "a \<in> {0..1}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3527
    shows "winding_number(shiftpath a \<gamma>) z = winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3528
  apply (rule winding_number_unique_loop)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3529
  using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3530
  apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3531
  apply (frule winding_number [OF \<gamma>], clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3532
  apply (rule_tac x="shiftpath a p" in exI)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3533
  apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3534
  apply (auto simp: shiftpath_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3535
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3536
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3537
lemma winding_number_split_linepath:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3538
  assumes "c \<in> closed_segment a b" "z \<notin> closed_segment a b"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3539
    shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3540
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3541
  have "z \<notin> closed_segment a c" "z \<notin> closed_segment c b"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3542
    using assms  apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3543
    using assms  by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3544
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3545
    using assms
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3546
    by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3547
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3548
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3549
lemma winding_number_cong:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3550
   "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p t = q t) \<Longrightarrow> winding_number p z = winding_number q z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3551
  by (simp add: winding_number_def pathstart_def pathfinish_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3552
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3553
lemma winding_number_offset: "winding_number p z = winding_number (\<lambda>w. p w - z) 0"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3554
  apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3555
  apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3556
  apply (rename_tac g)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3557
  apply (rule_tac x="\<lambda>t. g t - z" in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3558
  apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3559
  apply (rename_tac g)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3560
  apply (rule_tac x="\<lambda>t. g t + z" in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3561
  apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3562
  apply (force simp: algebra_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3563
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3564
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3565
(* A combined theorem deducing several things piecewise.*)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3566
lemma winding_number_join_pos_combined:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3567
     "\<lbrakk>valid_path \<gamma>1; z \<notin> path_image \<gamma>1; 0 < Re(winding_number \<gamma>1 z);
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3568
       valid_path \<gamma>2; z \<notin> path_image \<gamma>2; 0 < Re(winding_number \<gamma>2 z); pathfinish \<gamma>1 = pathstart \<gamma>2\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3569
      \<Longrightarrow> valid_path(\<gamma>1 +++ \<gamma>2) \<and> z \<notin> path_image(\<gamma>1 +++ \<gamma>2) \<and> 0 < Re(winding_number(\<gamma>1 +++ \<gamma>2) z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3570
  by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3571
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3572
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3573
(* Useful sufficient conditions for the winding number to be positive etc.*)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3574
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3575
lemma Re_winding_number:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3576
    "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3577
     \<Longrightarrow> Re(winding_number \<gamma> z) = Im(contour_integral \<gamma> (\<lambda>w. 1/(w - z))) / (2*pi)"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3578
by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3579
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3580
lemma winding_number_pos_le:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3581
  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3582
      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> 0 \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3583
    shows "0 \<le> Re(winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3584
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3585
  have *: "0 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" if x: "0 < x" "x < 1" for x
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3586
    using ge by (simp add: Complex.Im_divide algebra_simps x)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3587
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3588
    apply (simp add: Re_winding_number [OF \<gamma>] field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3589
    apply (rule has_integral_component_nonneg
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3590
             [of ii "\<lambda>x. if x \<in> {0<..<1}
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3591
                         then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else 0", simplified])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3592
      prefer 3 apply (force simp: *)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3593
     apply (simp add: Basis_complex_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3594
    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3595
    apply simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3596
    apply (simp only: box_real)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3597
    apply (subst has_contour_integral [symmetric])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3598
    using \<gamma>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3599
    apply (simp add: contour_integrable_inversediff has_contour_integral_integral)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3600
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3601
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3602
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3603
lemma winding_number_pos_lt_lemma:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3604
  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3605
      and e: "0 < e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3606
      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3607
    shows "0 < Re(winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3608
proof -
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3609
  have "e \<le> Im (contour_integral \<gamma> (\<lambda>w. 1 / (w - z)))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3610
    apply (rule has_integral_component_le
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3611
             [of ii "\<lambda>x. ii*e" "ii*e" "{0..1}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3612
                    "\<lambda>x. if x \<in> {0<..<1} then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else ii*e"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3613
                    "contour_integral \<gamma> (\<lambda>w. 1/(w - z))", simplified])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3614
    using e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3615
    apply (simp_all add: Basis_complex_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3616
    using has_integral_const_real [of _ 0 1] apply force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3617
    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)", simplified box_real])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3618
    apply simp
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3619
    apply (subst has_contour_integral [symmetric])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3620
    using \<gamma>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3621
    apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3622
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3623
  with e show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3624
    by (simp add: Re_winding_number [OF \<gamma>] field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3625
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3626
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3627
lemma winding_number_pos_lt:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3628
  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3629
      and e: "0 < e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3630
      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3631
    shows "0 < Re (winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3632
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3633
  have bm: "bounded ((\<lambda>w. w - z) ` (path_image \<gamma>))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3634
    using bounded_translation [of _ "-z"] \<gamma> by (simp add: bounded_valid_path_image)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3635
  then obtain B where B: "B > 0" and Bno: "\<And>x. x \<in> (\<lambda>w. w - z) ` (path_image \<gamma>) \<Longrightarrow> norm x \<le> B"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3636
    using bounded_pos [THEN iffD1, OF bm] by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3637
  { fix x::real  assume x: "0 < x" "x < 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3638
    then have B2: "cmod (\<gamma> x - z)^2 \<le> B^2" using Bno [of "\<gamma> x - z"]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3639
      by (simp add: path_image_def power2_eq_square mult_mono')
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3640
    with x have "\<gamma> x \<noteq> z" using \<gamma>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3641
      using path_image_def by fastforce
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3642
    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) * cnj (\<gamma> x - z)) / (cmod (\<gamma> x - z))\<^sup>2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3643
      using B ge [OF x] B2 e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3644
      apply (rule_tac y="e / (cmod (\<gamma> x - z))\<^sup>2" in order_trans)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3645
      apply (auto simp: divide_left_mono divide_right_mono)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3646
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3647
    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3648
      by (simp add: Im_divide_Reals complex_div_cnj [of _ "\<gamma> x - z" for x] del: complex_cnj_diff times_complex.sel)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3649
  } note * = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3650
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3651
    using e B by (simp add: * winding_number_pos_lt_lemma [OF \<gamma>, of "e/B^2"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3652
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3653
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3654
subsection\<open>The winding number is an integer\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3655
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3656
text\<open>Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1,
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3657
     Also on page 134 of Serge Lang's book with the name title, etc.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3658
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3659
lemma exp_fg:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3660
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3661
  assumes g: "(g has_vector_derivative g') (at x within s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3662
      and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3663
      and z: "g x \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3664
    shows "((\<lambda>x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3665
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3666
  have *: "(exp o (\<lambda>x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3667
    using assms unfolding has_vector_derivative_def scaleR_conv_of_real
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3668
    by (auto intro!: derivative_eq_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3669
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3670
    apply (rule has_vector_derivative_eq_rhs)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3671
    apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3672
    using z
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3673
    apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3674
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3675
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3676
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3677
lemma winding_number_exp_integral:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3678
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3679
  assumes \<gamma>: "\<gamma> piecewise_C1_differentiable_on {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3680
      and ab: "a \<le> b"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3681
      and z: "z \<notin> \<gamma> ` {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3682
    shows "(\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)) integrable_on {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3683
          (is "?thesis1")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3684
          "exp (- (integral {a..b} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))) * (\<gamma> b - z) = \<gamma> a - z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3685
          (is "?thesis2")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3686
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3687
  let ?D\<gamma> = "\<lambda>x. vector_derivative \<gamma> (at x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3688
  have [simp]: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<gamma> x \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3689
    using z by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3690
  have cong: "continuous_on {a..b} \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3691
    using \<gamma> by (simp add: piecewise_C1_differentiable_on_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3692
  obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3693
    using \<gamma> by (force simp: piecewise_C1_differentiable_on_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3694
  have o: "open ({a<..<b} - k)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3695
    using `finite k` by (simp add: finite_imp_closed open_Diff)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3696
  moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3697
    by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3698
  ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3699
    by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3700
  { fix w
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3701
    assume "w \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3702
    have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3703
      by (auto simp: dist_norm intro!: continuous_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3704
    moreover have "\<And>x. cmod (w - x) < cmod (w - z) \<Longrightarrow> \<exists>f'. ((\<lambda>w. 1 / (w - z)) has_field_derivative f') (at x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3705
      by (auto simp: intro!: derivative_eq_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3706
    ultimately have "\<exists>h. \<forall>y. norm(y - w) < norm(w - z) \<longrightarrow> (h has_field_derivative 1/(y - z)) (at y)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3707
      using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\<lambda>w. 1/(w - z)"]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3708
      by (simp add: complex_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3709
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3710
  then obtain h where h: "\<And>w y. w \<noteq> z \<Longrightarrow> norm(y - w) < norm(w - z) \<Longrightarrow> (h w has_field_derivative 1/(y - z)) (at y)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3711
    by meson
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3712
  have exy: "\<exists>y. ((\<lambda>x. inverse (\<gamma> x - z) * ?D\<gamma> x) has_integral y) {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3713
    unfolding integrable_on_def [symmetric]
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3714
    apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \<gamma>], of "-{z}"])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3715
    apply (rename_tac w)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3716
    apply (rule_tac x="norm(w - z)" in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3717
    apply (simp_all add: inverse_eq_divide)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3718
    apply (metis has_field_derivative_at_within h)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3719
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3720
  have vg_int: "(\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)) integrable_on {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3721
    unfolding box_real [symmetric] divide_inverse_commute
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3722
    by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3723
  with ab show ?thesis1
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3724
    by (simp add: divide_inverse_commute integral_def integrable_on_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3725
  { fix t
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3726
    assume t: "t \<in> {a..b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3727
    have cball: "continuous_on (ball (\<gamma> t) (dist (\<gamma> t) z)) (\<lambda>x. inverse (x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3728
        using z by (auto intro!: continuous_intros simp: dist_norm)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3729
    have icd: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> (\<lambda>w. inverse (w - z)) complex_differentiable at x"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3730
      unfolding complex_differentiable_def by (force simp: intro!: derivative_eq_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3731
    obtain h where h: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3732
                       (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\<gamma> t - y) < cmod (\<gamma> t - z)})"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3733
      using holomorphic_convex_primitive [where f = "\<lambda>w. inverse(w - z)", OF convex_ball finite.emptyI cball icd]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3734
      by simp (auto simp: ball_def dist_norm that)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3735
    { fix x D
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3736
      assume x: "x \<notin> k" "a < x" "x < b"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3737
      then have "x \<in> interior ({a..b} - k)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3738
        using open_subset_interior [OF o] by fastforce
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3739
      then have con: "isCont (\<lambda>x. ?D\<gamma> x) x"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3740
        using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3741
      then have con_vd: "continuous (at x within {a..b}) (\<lambda>x. ?D\<gamma> x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3742
        by (rule continuous_at_imp_continuous_within)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3743
      have gdx: "\<gamma> differentiable at x"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3744
        using x by (simp add: g_diff_at)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3745
      have "((\<lambda>c. exp (- integral {a..c} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))) * (\<gamma> c - z)) has_derivative (\<lambda>h. 0))
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3746
          (at x within {a..b})"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3747
        using x gdx t
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3748
        apply (clarsimp simp add: differentiable_iff_scaleR)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3749
        apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3750
        apply (simp_all add: has_vector_derivative_def [symmetric])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3751
        apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3752
        apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3753
        done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3754
      } note * = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3755
    have "exp (- (integral {a..t} (\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)))) * (\<gamma> t - z) =\<gamma> a - z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3756
      apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \<union> k" a b])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3757
      using t
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3758
      apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int]  simp add: ab)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3759
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3760
   }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3761
  with ab show ?thesis2
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3762
    by (simp add: divide_inverse_commute integral_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3763
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3764
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3765
corollary winding_number_exp_2pi:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3766
    "\<lbrakk>path p; z \<notin> path_image p\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3767
     \<Longrightarrow> pathfinish p - z = exp (2 * pi * ii * winding_number p z) * (pathstart p - z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3768
using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3769
  by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3770
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3771
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3772
subsection\<open>The version with complex integers and equality\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3773
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3774
lemma integer_winding_number_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3775
  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3776
  shows "winding_number \<gamma> z \<in> \<int> \<longleftrightarrow> pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3777
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3778
  have *: "\<And>i::complex. \<And>g0 g1. \<lbrakk>i \<noteq> 0; g0 \<noteq> z; (g1 - z) / i = g0 - z\<rbrakk> \<Longrightarrow> (i = 1 \<longleftrightarrow> g1 = g0)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3779
      by (simp add: field_simps) algebra
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3780
  obtain p where p: "valid_path p" "z \<notin> path_image p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3781
                    "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3782
                    "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3783
    using winding_number [OF assms, of 1] by auto
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3784
  have [simp]: "(winding_number \<gamma> z \<in> \<int>) = (exp (contour_integral p (\<lambda>w. 1 / (w - z))) = 1)"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3785
      using p by (simp add: exp_eq_1 complex_is_Int_iff)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3786
  have "winding_number p z \<in> \<int> \<longleftrightarrow> pathfinish p = pathstart p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3787
    using p z
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3788
    apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3789
    using winding_number_exp_integral(2) [of p 0 1 z]
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3790
    apply (simp add: field_simps contour_integral_integral exp_minus)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3791
    apply (rule *)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3792
    apply (auto simp: path_image_def field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3793
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3794
  then show ?thesis using p
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3795
    by (auto simp: winding_number_valid_path)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3796
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3797
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3798
theorem integer_winding_number:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3799
  "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> \<int>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3800
by (metis integer_winding_number_eq)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3801
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3802
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3803
text\<open>If the winding number's magnitude is at least one, then the path must contain points in every direction.*)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3804
   We can thus bound the winding number of a path that doesn't intersect a given ray. \<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3805
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3806
lemma winding_number_pos_meets:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3807
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3808
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and 1: "Re (winding_number \<gamma> z) \<ge> 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3809
      and w: "w \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3810
  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3811
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3812
  have [simp]: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<gamma> x \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3813
    using z by (auto simp: path_image_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3814
  have [simp]: "z \<notin> \<gamma> ` {0..1}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3815
    using path_image_def z by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3816
  have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3817
    using \<gamma> valid_path_def by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3818
  def r \<equiv> "(w - z) / (\<gamma> 0 - z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3819
  have [simp]: "r \<noteq> 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3820
    using w z by (auto simp: r_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3821
  have "Arg r \<le> 2*pi"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3822
    by (simp add: Arg less_eq_real_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3823
  also have "... \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3824
    using 1
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3825
    apply (simp add: winding_number_valid_path [OF \<gamma> z] Cauchy_Integral_Thm.contour_integral_integral)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3826
    apply (simp add: Complex.Re_divide field_simps power2_eq_square)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3827
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3828
  finally have "Arg r \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3829
  then have "\<exists>t. t \<in> {0..1} \<and> Im(integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3830
    apply (simp add:)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3831
    apply (rule Topological_Spaces.IVT')
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3832
    apply (simp_all add: Complex_Transcendental.Arg_ge_0)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3833
    apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3834
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3835
  then obtain t where t:     "t \<in> {0..1}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3836
                  and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3837
    by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3838
  def i \<equiv> "integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3839
  have iArg: "Arg r = Im i"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3840
    using eqArg by (simp add: i_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3841
  have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3842
    by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3843
  have "exp (- i) * (\<gamma> t - z) = \<gamma> 0 - z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3844
    unfolding i_def
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3845
    apply (rule winding_number_exp_integral [OF gpdt])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3846
    using t z unfolding path_image_def
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3847
    apply force+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3848
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3849
  then have *: "\<gamma> t - z = exp i * (\<gamma> 0 - z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3850
    by (simp add: exp_minus field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3851
  then have "(w - z) = r * (\<gamma> 0 - z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3852
    by (simp add: r_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3853
  then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \<gamma> t"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3854
    apply (simp add:)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3855
    apply (subst Complex_Transcendental.Arg_eq [of r])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3856
    apply (simp add: iArg)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3857
    using *
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  3858
    apply (simp add: exp_eq_polar field_simps)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3859
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3860
  with t show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3861
    by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3862
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3863
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3864
lemma winding_number_big_meets:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3865
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3866
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "abs (Re (winding_number \<gamma> z)) \<ge> 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3867
      and w: "w \<noteq> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3868
  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3869
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3870
  { assume "Re (winding_number \<gamma> z) \<le> - 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3871
    then have "Re (winding_number (reversepath \<gamma>) z) \<ge> 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3872
      by (simp add: \<gamma> valid_path_imp_path winding_number_reversepath z)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3873
    moreover have "valid_path (reversepath \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3874
      using \<gamma> valid_path_imp_reverse by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3875
    moreover have "z \<notin> path_image (reversepath \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3876
      by (simp add: z)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3877
    ultimately have "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image (reversepath \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3878
      using winding_number_pos_meets w by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3879
    then have ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3880
      by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3881
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3882
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3883
    using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3884
    by (simp add: Groups.abs_if_class.abs_if winding_number_pos_meets split: split_if_asm)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3885
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3886
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3887
lemma winding_number_less_1:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3888
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3889
  shows
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3890
  "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>; w \<noteq> z;
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3891
    \<And>a::real. 0 < a \<Longrightarrow> z + a*(w - z) \<notin> path_image \<gamma>\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3892
   \<Longrightarrow> abs (Re(winding_number \<gamma> z)) < 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3893
   by (auto simp: not_less dest: winding_number_big_meets)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3894
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3895
text\<open>One way of proving that WN=1 for a loop.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3896
lemma winding_number_eq_1:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3897
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3898
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3899
      and 0: "0 < Re(winding_number \<gamma> z)" and 2: "Re(winding_number \<gamma> z) < 2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3900
  shows "winding_number \<gamma> z = 1"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3901
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3902
  have "winding_number \<gamma> z \<in> Ints"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3903
    by (simp add: \<gamma> integer_winding_number loop valid_path_imp_path z)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3904
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3905
    using 0 2 by (auto simp: Ints_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3906
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3907
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3908
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3909
subsection\<open>Continuity of winding number and invariance on connected sets.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3910
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3911
lemma continuous_at_winding_number:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3912
  fixes z::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3913
  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3914
  shows "continuous (at z) (winding_number \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3915
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3916
  obtain e where "e>0" and cbg: "cball z e \<subseteq> - path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3917
    using open_contains_cball [of "- path_image \<gamma>"]  z
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3918
    by (force simp: closed_def [symmetric] closed_path_image [OF \<gamma>])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3919
  then have ppag: "path_image \<gamma> \<subseteq> - cball z (e/2)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3920
    by (force simp: cball_def dist_norm)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3921
  have oc: "open (- cball z (e / 2))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3922
    by (simp add: closed_def [symmetric])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3923
  obtain d where "d>0" and pi_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3924
    "\<And>h1 h2. \<lbrakk>valid_path h1; valid_path h2;
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3925
              (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d);
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3926
              pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3927
             \<Longrightarrow>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3928
               path_image h1 \<subseteq> - cball z (e / 2) \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3929
               path_image h2 \<subseteq> - cball z (e / 2) \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3930
               (\<forall>f. f holomorphic_on - cball z (e / 2) \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3931
    using contour_integral_nearby_ends [OF oc \<gamma> ppag] by metis
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3932
  obtain p where p: "valid_path p" "z \<notin> path_image p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3933
                    "pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3934
              and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3935
              and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3936
    using winding_number [OF \<gamma> z, of "min d e / 2"] `d>0` `e>0` by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3937
  { fix w
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3938
    assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3939
    then have wnotp: "w \<notin> path_image p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3940
      using cbg `d>0` `e>0`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3941
      apply (simp add: path_image_def cball_def dist_norm, clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3942
      apply (frule pg)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3943
      apply (drule_tac c="\<gamma> x" in subsetD)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3944
      apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3945
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3946
    have wnotg: "w \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3947
      using cbg e2 `e>0` by (force simp: dist_norm norm_minus_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3948
    { fix k::real
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3949
      assume k: "k>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3950
      then obtain q where q: "valid_path q" "w \<notin> path_image q"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3951
                             "pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3952
                    and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3953
                    and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3954
        using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] `d>0` `e>0` k
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3955
        by (force simp: min_divide_distrib_right)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3956
      have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>u. 1 / (u - w))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3957
        apply (rule pi_eq [OF `valid_path q` `valid_path p`, THEN conjunct2, THEN conjunct2, rule_format])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3958
        apply (frule pg)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3959
        apply (frule qg)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3960
        using p q `d>0` e2
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3961
        apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3962
        done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3963
      then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3964
        by (simp add: pi qi)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3965
    } note pip = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3966
    have "path p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3967
      using p by (simp add: valid_path_imp_path)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3968
    then have "winding_number p w = winding_number \<gamma> w"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3969
      apply (rule winding_number_unique [OF _ wnotp])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3970
      apply (rule_tac x=p in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3971
      apply (simp add: p wnotp min_divide_distrib_right pip)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3972
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3973
  } note wnwn = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3974
  obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \<subseteq> - path_image p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3975
    using p open_contains_cball [of "- path_image p"]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3976
    by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3977
  obtain L
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3978
    where "L>0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3979
      and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe);
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3980
                      \<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3981
                      cmod (contour_integral p f) \<le> L * B"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3982
    using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp `valid_path p` by force
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3983
  { fix e::real and w::complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3984
    assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3985
    then have [simp]: "w \<notin> path_image p"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3986
      using cbp p(2) `0 < pe`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3987
      by (force simp: dist_norm norm_minus_commute path_image_def cball_def)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3988
    have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) =
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3989
                  contour_integral p (\<lambda>x. 1/(x - w) - 1/(x - z))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3990
      by (simp add: p contour_integrable_inversediff contour_integral_diff)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3991
    { fix x
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3992
      assume pe: "3/4 * pe < cmod (z - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3993
      have "cmod (w - x) < pe/4 + cmod (z - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3994
        by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1))
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3995
      then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3996
      have "cmod (z - x) \<le> cmod (z - w) + cmod (w - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3997
        using norm_diff_triangle_le by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3998
      also have "... < pe/4 + cmod (w - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  3999
        using w by (simp add: norm_minus_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4000
      finally have "pe/2 < cmod (w - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4001
        using pe by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4002
      then have "(pe/2)^2 < cmod (w - x) ^ 2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4003
        apply (rule power_strict_mono)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4004
        using `pe>0` by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4005
      then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
  4006
        by (simp add: power_divide)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4007
      have "8 * L * cmod (w - z) < e * pe\<^sup>2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4008
        using w `L>0` by (simp add: field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4009
      also have "... < e * 4 * cmod (w - x) * cmod (w - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4010
        using pe2 `e>0` by (simp add: power2_eq_square)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4011
      also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4012
        using wx
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4013
        apply (rule mult_strict_left_mono)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4014
        using pe2 e not_less_iff_gr_or_eq by fastforce
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4015
      finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4016
        by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4017
      also have "... \<le> e * cmod (w - x) * cmod (z - x)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4018
         using e by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4019
      finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4020
      have "L * cmod (1 / (x - w) - 1 / (x - z)) \<le> e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4021
        apply (cases "x=z \<or> x=w")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4022
        using pe `pe>0` w `L>0`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4023
        apply (force simp: norm_minus_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4024
        using wx w(2) `L>0` pe pe2 Lwz
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4025
        apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4026
        done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4027
    } note L_cmod_le = this
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4028
    have *: "cmod (contour_integral p (\<lambda>x. 1 / (x - w) - 1 / (x - z))) \<le> L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4029
      apply (rule L)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4030
      using `pe>0` w
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4031
      apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4032
      using `pe>0` w `L>0`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4033
      apply (auto simp: cball_def dist_norm field_simps L_cmod_le  simp del: less_divide_eq_numeral1 le_divide_eq_numeral1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4034
      done
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4035
    have "cmod (contour_integral p (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4036
      apply (simp add:)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4037
      apply (rule le_less_trans [OF *])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4038
      using `L>0` e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4039
      apply (force simp: field_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4040
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4041
    then have "cmod (winding_number p w - winding_number p z) < e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4042
      using pi_ge_two e
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4043
      by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4044
  } note cmod_wn_diff = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4045
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4046
    apply (rule continuous_transform_at [of "min d e / 2" _ "winding_number p"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4047
    apply (auto simp: `d>0` `e>0` dist_norm wnwn simp del: less_divide_eq_numeral1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4048
    apply (simp add: continuous_at_eps_delta, clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4049
    apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4050
    using `pe>0` `L>0`
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4051
    apply (simp add: dist_norm cmod_wn_diff)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4052
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4053
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4054
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4055
corollary continuous_on_winding_number:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4056
    "path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4057
  by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4058
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4059
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4060
subsection{*The winding number is constant on a connected region*}
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4061
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4062
lemma winding_number_constant:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4063
  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4064
    shows "\<exists>k. \<forall>z \<in> s. winding_number \<gamma> z = k"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4065
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4066
  { fix y z
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4067
    assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4068
    assume "y \<in> s" "z \<in> s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4069
    then have "winding_number \<gamma> y \<in> \<int>"  "winding_number \<gamma> z \<in>  \<int>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4070
      using integer_winding_number [OF \<gamma> loop] sg `y \<in> s` by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4071
    with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4072
      by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4073
  } note * = this
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4074
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4075
    apply (rule continuous_discrete_range_constant [OF cs])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4076
    using continuous_on_winding_number [OF \<gamma>] sg
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4077
    apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4078
    apply (rule_tac x=1 in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4079
    apply (auto simp: *)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4080
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4081
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4082
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4083
lemma winding_number_eq:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4084
     "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; w \<in> s; z \<in> s; connected s; s \<inter> path_image \<gamma> = {}\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4085
      \<Longrightarrow> winding_number \<gamma> w = winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4086
using winding_number_constant by fastforce
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4087
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4088
lemma open_winding_number_levelsets:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4089
  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4090
    shows "open {z. z \<notin> path_image \<gamma> \<and> winding_number \<gamma> z = k}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4091
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4092
  have op: "open (- path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4093
    by (simp add: closed_path_image \<gamma> open_Compl)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4094
  { fix z assume z: "z \<notin> path_image \<gamma>" and k: "k = winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4095
    obtain e where e: "e>0" "ball z e \<subseteq> - path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4096
      using open_contains_ball [of "- path_image \<gamma>"] op z
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4097
      by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4098
    have "\<exists>e>0. \<forall>y. dist y z < e \<longrightarrow> y \<notin> path_image \<gamma> \<and> winding_number \<gamma> y = winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4099
      apply (rule_tac x=e in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4100
      using e apply (simp add: dist_norm ball_def norm_minus_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4101
      apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4102
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4103
  } then
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4104
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4105
    by (auto simp: Complex.open_complex_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4106
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4107
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4108
subsection\<open>Winding number is zero "outside" a curve, in various senses\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4109
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4110
lemma winding_number_zero_in_outside:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4111
  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4112
    shows "winding_number \<gamma> z = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4113
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4114
  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4115
    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4116
  obtain w::complex where w: "w \<notin> ball 0 (B + 1)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4117
    by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4118
  have "- ball 0 (B + 1) \<subseteq> outside (path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4119
    apply (rule outside_subset_convex)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4120
    using B subset_ball by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4121
  then have wout: "w \<in> outside (path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4122
    using w by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4123
  moreover obtain k where "\<And>z. z \<in> outside (path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = k"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4124
    using winding_number_constant [OF \<gamma> loop, of "outside(path_image \<gamma>)"] connected_outside
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4125
    by (metis DIM_complex bounded_path_image dual_order.refl \<gamma> outside_no_overlap)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4126
  ultimately have "winding_number \<gamma> z = winding_number \<gamma> w"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4127
    using z by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4128
  also have "... = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4129
  proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4130
    have wnot: "w \<notin> path_image \<gamma>"  using wout by (simp add: outside_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4131
    { fix e::real assume "0<e"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4132
      obtain p where p: "polynomial_function p" "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4133
                 and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4134
                 and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4135
        using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] `e>0` by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4136
      have pip: "path_image p \<subseteq> ball 0 (B + 1)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4137
        using B
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4138
        apply (clarsimp simp add: path_image_def dist_norm ball_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4139
        apply (frule (1) pg1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4140
        apply (fastforce dest: norm_add_less)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4141
        done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4142
      then have "w \<notin> path_image p"  using w by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4143
      then have "\<exists>p. valid_path p \<and> w \<notin> path_image p \<and>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4144
                     pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4145
                     (\<forall>t\<in>{0..1}. cmod (\<gamma> t - p t) < e) \<and> contour_integral p (\<lambda>wa. 1 / (wa - w)) = 0"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4146
        apply (rule_tac x=p in exI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4147
        apply (simp add: p valid_path_polynomial_function)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4148
        apply (intro conjI)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4149
        using pge apply (simp add: norm_minus_commute)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4150
        apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4151
        apply (rule holomorphic_intros | simp add: dist_norm)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4152
        using mem_ball_0 w apply blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4153
        using p apply (simp_all add: valid_path_polynomial_function loop pip)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4154
        done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4155
    }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4156
    then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4157
      by (auto intro: winding_number_unique [OF \<gamma>] simp add: wnot)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4158
  qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4159
  finally show ?thesis .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4160
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4161
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4162
lemma winding_number_zero_outside:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4163
    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4164
  by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4165
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4166
lemma winding_number_zero_at_infinity:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4167
  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4168
    shows "\<exists>B. \<forall>z. B \<le> norm z \<longrightarrow> winding_number \<gamma> z = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4169
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4170
  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4171
    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4172
  then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4173
    apply (rule_tac x="B+1" in exI, clarify)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4174
    apply (rule winding_number_zero_outside [OF \<gamma> convex_cball [of 0 B] loop])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4175
    apply (meson less_add_one mem_cball_0 not_le order_trans)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4176
    using ball_subset_cball by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4177
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4178
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4179
lemma winding_number_zero_point:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4180
    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; open s; path_image \<gamma> \<subseteq> s\<rbrakk>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4181
     \<Longrightarrow> \<exists>z. z \<in> s \<and> winding_number \<gamma> z = 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4182
  using outside_compact_in_open [of "path_image \<gamma>" s] path_image_nonempty winding_number_zero_in_outside
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4183
  by (fastforce simp add: compact_path_image)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4184
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4185
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4186
text\<open>If a path winds round a set, it winds rounds its inside.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4187
lemma winding_number_around_inside:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4188
  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4189
      and cls: "closed s" and cos: "connected s" and s_disj: "s \<inter> path_image \<gamma> = {}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4190
      and z: "z \<in> s" and wn_nz: "winding_number \<gamma> z \<noteq> 0" and w: "w \<in> s \<union> inside s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4191
    shows "winding_number \<gamma> w = winding_number \<gamma> z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4192
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4193
  have ssb: "s \<subseteq> inside(path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4194
  proof
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4195
    fix x :: complex
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4196
    assume "x \<in> s"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4197
    hence "x \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4198
      by (meson disjoint_iff_not_equal s_disj)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4199
    thus "x \<in> inside (path_image \<gamma>)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4200
      using `x \<in> s` by (metis (no_types) ComplI UnE cos \<gamma> loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4201
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4202
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4203
    apply (rule winding_number_eq [OF \<gamma> loop w])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4204
    using z apply blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4205
    apply (simp add: cls connected_with_inside cos)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4206
    apply (simp add: Int_Un_distrib2 s_disj, safe)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4207
    by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \<gamma>"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \<gamma> inside_no_overlap)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4208
 qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4209
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4210
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4211
text\<open>Bounding a WN by 1/2 for a path and point in opposite halfspaces.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4212
lemma winding_number_subpath_continuous:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4213
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4214
    shows "continuous_on {0..1} (\<lambda>x. winding_number(subpath 0 x \<gamma>) z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4215
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4216
  have *: "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4217
         winding_number (subpath 0 x \<gamma>) z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4218
         if x: "0 \<le> x" "x \<le> 1" for x
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4219
  proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4220
    have "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4221
          1 / (2*pi*ii) * contour_integral (subpath 0 x \<gamma>) (\<lambda>w. 1/(w - z))"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4222
      using assms x
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4223
      apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4224
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4225
    also have "... = winding_number (subpath 0 x \<gamma>) z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4226
      apply (subst winding_number_valid_path)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4227
      using assms x
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  4228
      apply (simp_all add: path_image_subpath valid_path_subpath)
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  4229
      by (force simp: path_image_def)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4230
    finally show ?thesis .
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4231
  qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4232
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4233
    apply (rule continuous_on_eq
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4234
                 [where f = "\<lambda>x. 1 / (2*pi*ii) *
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4235
                                 integral {0..x} (\<lambda>t. 1/(\<gamma> t - z) * vector_derivative \<gamma> (at t))"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4236
    apply (rule continuous_intros)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4237
    apply (rule indefinite_integral_continuous)
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4238
    apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on])
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4239
      using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4240
    apply (simp add: *)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4241
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4242
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4243
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4244
lemma winding_number_ivt_pos:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4245
    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> Re(winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4246
      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4247
  apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4248
  apply (simp add:)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4249
  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4250
  using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4251
  apply (auto simp: path_image_def image_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4252
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4253
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4254
lemma winding_number_ivt_neg:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4255
    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "Re(winding_number \<gamma> z) \<le> w" "w \<le> 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4256
      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4257
  apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4258
  apply (simp add:)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4259
  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4260
  using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4261
  apply (auto simp: path_image_def image_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4262
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4263
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4264
lemma winding_number_ivt_abs:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4265
    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> \<bar>Re(winding_number \<gamma> z)\<bar>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4266
      shows "\<exists>t \<in> {0..1}. \<bar>Re (winding_number (subpath 0 t \<gamma>) z)\<bar> = w"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4267
  using assms winding_number_ivt_pos [of \<gamma> z w] winding_number_ivt_neg [of \<gamma> z "-w"]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4268
  by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4269
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4270
lemma winding_number_lt_half_lemma:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4271
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and az: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4272
    shows "Re(winding_number \<gamma> z) < 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4273
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4274
  { assume "Re(winding_number \<gamma> z) \<ge> 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4275
    then obtain t::real where t: "0 \<le> t" "t \<le> 1" and sub12: "Re (winding_number (subpath 0 t \<gamma>) z) = 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4276
      using winding_number_ivt_pos [OF \<gamma> z, of "1/2"] by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4277
    have gt: "\<gamma> t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \<gamma>) z)))) * (\<gamma> 0 - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4278
      using winding_number_exp_2pi [of "subpath 0 t \<gamma>" z]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4279
      apply (simp add: t \<gamma> valid_path_imp_path)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  4280
      using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12)
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4281
    have "b < a \<bullet> \<gamma> 0"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4282
    proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4283
      have "\<gamma> 0 \<in> {c. b < a \<bullet> c}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4284
        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4285
      thus ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4286
        by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4287
    qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4288
    moreover have "b < a \<bullet> \<gamma> t"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4289
    proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4290
      have "\<gamma> t \<in> {c. b < a \<bullet> c}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4291
        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4292
      thus ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4293
        by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4294
    qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4295
    ultimately have "0 < a \<bullet> (\<gamma> 0 - z)" "0 < a \<bullet> (\<gamma> t - z)" using az
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4296
      by (simp add: inner_diff_right)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4297
    then have False
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4298
      by (simp add: gt inner_mult_right mult_less_0_iff)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4299
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4300
  then show ?thesis by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4301
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4302
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4303
lemma winding_number_lt_half:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4304
  assumes "valid_path \<gamma>" "a \<bullet> z \<le> b" "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4305
    shows "\<bar>Re (winding_number \<gamma> z)\<bar> < 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4306
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4307
  have "z \<notin> path_image \<gamma>" using assms by auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4308
  with assms show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4309
    apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4310
    apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \<gamma> z a b]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4311
                 winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4312
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4313
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4314
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4315
lemma winding_number_le_half:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4316
  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4317
      and anz: "a \<noteq> 0" and azb: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w \<ge> b}"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4318
    shows "\<bar>Re (winding_number \<gamma> z)\<bar> \<le> 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4319
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4320
  { assume wnz_12: "\<bar>Re (winding_number \<gamma> z)\<bar> > 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4321
    have "isCont (winding_number \<gamma>) z"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4322
      by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4323
    then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < abs(Re(winding_number \<gamma> z)) - 1/2"
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61738
diff changeset
  4324
      using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4325
    def z' \<equiv> "z - (d / (2 * cmod a)) *\<^sub>R a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4326
    have *: "a \<bullet> z' \<le> b - d / 3 * cmod a"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4327
      unfolding z'_def inner_mult_right' divide_inverse
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4328
      apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4329
      apply (metis `0 < d` add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4330
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4331
    have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4332
      using d [of z'] anz `d>0` by (simp add: dist_norm z'_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4333
    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4334
      by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4335
    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4336
      using abs_Re_le_cmod [of "winding_number \<gamma> z' - winding_number \<gamma> z"] by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4337
    then have wnz_12': "\<bar>Re (winding_number \<gamma> z')\<bar> > 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4338
      by linarith
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4339
    moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4340
      apply (rule winding_number_lt_half [OF \<gamma> *])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4341
      using azb `d>0` pag
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4342
      apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4343
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4344
    ultimately have False
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4345
      by simp
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4346
  }
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4347
  then show ?thesis by force
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4348
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4349
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4350
lemma winding_number_lt_half_linepath: "z \<notin> closed_segment a b \<Longrightarrow> \<bar>Re (winding_number (linepath a b) z)\<bar> < 1/2"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4351
  using separating_hyperplane_closed_point [of "closed_segment a b" z]
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4352
  apply auto
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4353
  apply (simp add: closed_segment_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4354
  apply (drule less_imp_le)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4355
  apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4356
  apply (auto simp: segment)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4357
  done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4358
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4359
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4360
text\<open> Positivity of WN for a linepath.\<close>
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4361
lemma winding_number_linepath_pos_lt:
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4362
    assumes "0 < Im ((b - a) * cnj (b - z))"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4363
      shows "0 < Re(winding_number(linepath a b) z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4364
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4365
  have z: "z \<notin> path_image (linepath a b)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4366
    using assms
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4367
    by (simp add: closed_segment_def) (force simp: algebra_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4368
  show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4369
    apply (rule winding_number_pos_lt [OF valid_path_linepath z assms])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4370
    apply (simp add: linepath_def algebra_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4371
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4372
qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4373
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4374
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4375
subsection{* Cauchy's integral formula, again for a convex enclosing set.*}
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4376
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: 61520
diff changeset
  4377
lemma Cauchy_integral_formula_weak:
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: 61520
diff changeset
  4378
    assumes s: "convex s" and "finite k" and conf: "continuous_on s f"
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: 61520
diff changeset
  4379
        and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f complex_differentiable at x)"
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: 61520
diff changeset
  4380
        and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4381
        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4382
      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * ii * winding_number \<gamma> z * f z)) \<gamma>"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4383
proof -
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4384
  obtain f' where f': "(f has_field_derivative f') (at z)"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4385
    using fcd [OF z] by (auto simp: complex_differentiable_def)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4386
  have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4387
  have c: "continuous (at x within s) (\<lambda>w. if w = z then f' else (f w - f z) / (w - z))" if "x \<in> s" for x
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4388
  proof (cases "x = z")
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4389
    case True then show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4390
      apply (simp add: continuous_within)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4391
      apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"])
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: 61520
diff changeset
  4392
      using has_field_derivative_at_within DERIV_within_iff f'
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4393
      apply (fastforce simp add:)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4394
      done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4395
  next
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4396
    case False
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4397
    then have dxz: "dist x z > 0" using dist_nz by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4398
    have cf: "continuous (at x within s) f"
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4399
      using conf continuous_on_eq_continuous_within that by blast
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4400
    show ?thesis
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4401
      apply (rule continuous_transform_within [OF dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4402
      apply (force simp: dist_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4403
      apply (rule cf continuous_intros)+
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: 61520
diff changeset
  4404
      using False by auto
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4405
  qed
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4406
  have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4407
  have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4408
    apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4409
    using c apply (force simp: continuous_on_eq_continuous_within)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4410
    apply (rename_tac w)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4411
    apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in complex_differentiable_transform_within)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4412
    apply (simp_all add: dist_pos_lt dist_commute)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4413
    apply (metis less_irrefl)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4414
    apply (rule derivative_intros fcd | simp)+
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4415
    done
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4416
  show ?thesis
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4417
    apply (rule has_contour_integral_eq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4418
    using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4419
    apply (auto simp: mult_ac divide_simps)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4420
    done
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: 61520
diff changeset
  4421
qed
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: 61520
diff changeset
  4422
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: 61520
diff changeset
  4423
theorem Cauchy_integral_formula_convex_simple:
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4424
    "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z};
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4425
      pathfinish \<gamma> = pathstart \<gamma>\<rbrakk>
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4426
     \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * ii * winding_number \<gamma> z * f z)) \<gamma>"
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4427
  apply (rule Cauchy_integral_formula_weak [where k = "{}"])
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4428
  using holomorphic_on_imp_continuous_on
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4429
  by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE)
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61518
diff changeset
  4430
61738
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4431
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4432
subsection\<open>Homotopy forms of Cauchy's theorem\<close>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4433
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4434
proposition Cauchy_theorem_homotopic:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4435
    assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4436
        and "open s" and f: "f holomorphic_on s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4437
        and vpg: "valid_path g" and vph: "valid_path h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4438
    shows "contour_integral g f = contour_integral h f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4439
proof -
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4440
  have pathsf: "linked_paths atends g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4441
    using hom  by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4442
  obtain k :: "real \<times> real \<Rightarrow> complex"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4443
    where contk: "continuous_on ({0..1} \<times> {0..1}) k"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4444
      and ks: "k ` ({0..1} \<times> {0..1}) \<subseteq> s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4445
      and k [simp]: "\<forall>x. k (0, x) = g x" "\<forall>x. k (1, x) = h x"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4446
      and ksf: "\<forall>t\<in>{0..1}. linked_paths atends g (\<lambda>x. k (t, x))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4447
      using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: split_if_asm)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4448
  have ucontk: "uniformly_continuous_on ({0..1} \<times> {0..1}) k"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4449
    by (blast intro: compact_Times compact_uniformly_continuous [OF contk])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4450
  { fix t::real assume t: "t \<in> {0..1}"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4451
    have pak: "path (k o (\<lambda>u. (t, u)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4452
      unfolding path_def
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4453
      apply (rule continuous_intros continuous_on_subset [OF contk])+
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4454
      using t by force
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4455
    have pik: "path_image (k \<circ> Pair t) \<subseteq> s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4456
      using ks t by (auto simp: path_image_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4457
    obtain e where "e>0" and e:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4458
         "\<And>g h. \<lbrakk>valid_path g; valid_path h;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4459
                  \<forall>u\<in>{0..1}. cmod (g u - (k \<circ> Pair t) u) < e \<and> cmod (h u - (k \<circ> Pair t) u) < e;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4460
                  linked_paths atends g h\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4461
                 \<Longrightarrow> contour_integral h f = contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4462
      using contour_integral_nearby [OF \<open>open s\<close> pak pik, of atends] f by metis
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4463
    obtain d where "d>0" and d:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4464
        "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> {0..1}; x' \<in> {0..1} \<times> {0..1}; norm (x'-x) < d\<rbrakk> \<Longrightarrow> norm (k x' - k x) < e/4"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4465
      by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm `e>0`)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4466
    { fix t1 t2
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4467
      assume t1: "0 \<le> t1" "t1 \<le> 1" and t2: "0 \<le> t2" "t2 \<le> 1" and ltd: "\<bar>t1 - t\<bar> < d" "\<bar>t2 - t\<bar> < d"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4468
      have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4469
        using \<open>e > 0\<close>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4470
        apply (rule_tac y = k1 in norm_triangle_half_l)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4471
        apply (auto simp: norm_minus_commute intro: order_less_trans)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4472
        done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4473
      have "\<exists>d>0. \<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4474
                          (\<forall>u\<in>{0..1}. cmod (g1 u - k (t1, u)) < d \<and> cmod (g2 u - k (t2, u)) < d) \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4475
                          linked_paths atends g1 g2 \<longrightarrow>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4476
                          contour_integral g2 f = contour_integral g1 f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4477
        apply (rule_tac x="e/4" in exI)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4478
        using t t1 t2 ltd \<open>e > 0\<close>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4479
        apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4480
        done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4481
    }
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4482
    then have "\<exists>e. 0 < e \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4483
              (\<forall>t1 t2. t1 \<in> {0..1} \<and> t2 \<in> {0..1} \<and> abs(t1 - t) < e \<and> abs(t2 - t) < e
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4484
                \<longrightarrow> (\<exists>d. 0 < d \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4485
                     (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4486
                       (\<forall>u \<in> {0..1}.
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4487
                          norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4488
                          linked_paths atends g1 g2
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4489
                          \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4490
      by (rule_tac x=d in exI) (simp add: \<open>d > 0\<close>)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4491
  }
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4492
  then obtain ee where ee:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4493
       "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0 \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4494
          (\<forall>t1 t2. t1 \<in> {0..1} \<longrightarrow> t2 \<in> {0..1} \<longrightarrow> abs(t1 - t) < ee t \<longrightarrow> abs(t2 - t) < ee t
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4495
            \<longrightarrow> (\<exists>d. 0 < d \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4496
                 (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4497
                   (\<forall>u \<in> {0..1}.
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4498
                      norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4499
                      linked_paths atends g1 g2
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4500
                      \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4501
    by metis
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4502
  note ee_rule = ee [THEN conjunct2, rule_format]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4503
  def C \<equiv> "(\<lambda>t. ball t (ee t / 3)) ` {0..1}"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4504
  have "\<forall>t \<in> C. open t" by (simp add: C_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4505
  moreover have "{0..1} \<subseteq> \<Union>C"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4506
    using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4507
  ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4508
    by (rule compactE [OF compact_interval])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4509
  def kk \<equiv> "{t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4510
  have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4511
  def e \<equiv> "Min (ee ` kk)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4512
  have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4513
    using C' by (auto simp: kk_def C_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4514
  have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4515
    by (simp add: kk_def ee)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4516
  moreover have "finite kk"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4517
    using \<open>finite C'\<close> kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4518
  moreover have "kk \<noteq> {}" using \<open>{0..1} \<subseteq> \<Union>C'\<close> C'_eq by force
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4519
  ultimately have "e > 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4520
    using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4521
  then obtain N::nat where "N > 0" and N: "1/N < e/3"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4522
    by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4523
  have e_le_ee: "\<And>i. i \<in> kk \<Longrightarrow> e \<le> ee i"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4524
    using \<open>finite kk\<close> by (simp add: e_def Min_le_iff [of "ee ` kk"])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4525
  have plus: "\<exists>t \<in> kk. x \<in> ball t (ee t / 3)" if "x \<in> {0..1}" for x
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4526
    using C' subsetD [OF C'01 that]  unfolding C'_eq by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4527
  have [OF order_refl]:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4528
      "\<exists>d. 0 < d \<and> (\<forall>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (n/N, u)) < d) \<and> linked_paths atends g j
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4529
                        \<longrightarrow> contour_integral j f = contour_integral g f)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4530
       if "n \<le> N" for n
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4531
  using that
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4532
  proof (induct n)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4533
    case 0 show ?case using ee_rule [of 0 0 0]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4534
      apply clarsimp
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4535
      apply (rule_tac x=d in exI, safe)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4536
      by (metis diff_self vpg norm_zero)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4537
  next
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4538
    case (Suc n)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4539
    then have N01: "n/N \<in> {0..1}" "(Suc n)/N \<in> {0..1}"  by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4540
    then obtain t where t: "t \<in> kk" "n/N \<in> ball t (ee t / 3)"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4541
      using plus [of "n/N"] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4542
    then have nN_less: "\<bar>n/N - t\<bar> < ee t"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4543
      by (simp add: dist_norm del: less_divide_eq_numeral1)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4544
    have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4545
      using t N \<open>N > 0\<close> e_le_ee [of t]
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4546
      by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4547
    have t01: "t \<in> {0..1}" using `kk \<subseteq> {0..1}` `t \<in> kk` by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4548
    obtain d1 where "d1 > 0" and d1:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4549
        "\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4550
                   \<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4551
                   linked_paths atends g1 g2\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4552
                   \<Longrightarrow> contour_integral g2 f = contour_integral g1 f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4553
      using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4554
    have "n \<le> N" using Suc.prems by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4555
    with Suc.hyps
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4556
    obtain d2 where "d2 > 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4557
      and d2: "\<And>j. \<lbrakk>valid_path j; \<forall>u\<in>{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4558
                     \<Longrightarrow> contour_integral j f = contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4559
        by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4560
    have "continuous_on {0..1} (k o (\<lambda>u. (n/N, u)))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4561
      apply (rule continuous_intros continuous_on_subset [OF contk])+
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4562
      using N01 by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4563
    then have pkn: "path (\<lambda>u. k (n/N, u))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4564
      by (simp add: path_def)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4565
    have min12: "min d1 d2 > 0" by (simp add: `0 < d1` `0 < d2`)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4566
    obtain p where "polynomial_function p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4567
        and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4568
                 "pathfinish p = pathfinish (\<lambda>u. k (n/N, u))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4569
        and pk_le:  "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (p t - k (n/N, t)) < min d1 d2"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4570
      using path_approx_polynomial_function [OF pkn min12] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4571
    then have vpp: "valid_path p" using valid_path_polynomial_function by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4572
    have lpa: "linked_paths atends g p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4573
      by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4574
    show ?case
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4575
      apply (rule_tac x="min d1 d2" in exI)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4576
      apply (simp add: `0 < d1` `0 < d2`, clarify)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4577
      apply (rule_tac s="contour_integral p f" in trans)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4578
      using pk_le N01(1) ksf pathfinish_def pathstart_def
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4579
      apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4580
      using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4581
      done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4582
  qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4583
  then obtain d where "0 < d"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4584
                       "\<And>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (1,u)) < d) \<and>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4585
                            linked_paths atends g j
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4586
                            \<Longrightarrow> contour_integral j f = contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4587
    using \<open>N>0\<close> by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4588
  then have "linked_paths atends g h \<Longrightarrow> contour_integral h f = contour_integral g f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4589
    using \<open>N>0\<close> vph by fastforce
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4590
  then show ?thesis
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4591
    by (simp add: pathsf)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4592
qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4593
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4594
proposition Cauchy_theorem_homotopic_paths:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4595
    assumes hom: "homotopic_paths s g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4596
        and "open s" and f: "f holomorphic_on s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4597
        and vpg: "valid_path g" and vph: "valid_path h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4598
    shows "contour_integral g f = contour_integral h f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4599
  using Cauchy_theorem_homotopic [of True s g h] assms by simp
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4600
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4601
proposition Cauchy_theorem_homotopic_loops:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4602
    assumes hom: "homotopic_loops s g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4603
        and "open s" and f: "f holomorphic_on s"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4604
        and vpg: "valid_path g" and vph: "valid_path h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4605
    shows "contour_integral g f = contour_integral h f"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4606
  using Cauchy_theorem_homotopic [of False s g h] assms by simp
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4607
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4608
lemma has_contour_integral_newpath:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4609
    "\<lbrakk>(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4610
     \<Longrightarrow> (f has_contour_integral y) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4611
  using has_contour_integral_integral contour_integral_unique by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4612
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4613
lemma Cauchy_theorem_null_homotopic:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4614
     "\<lbrakk>f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4615
  apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4616
  using contour_integrable_holomorphic_simple
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4617
    apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4618
  by (simp add: Cauchy_theorem_homotopic_loops)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4619
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4620
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4621
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4622
subsection\<open>More winding number properties\<close>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4623
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4624
text\<open>including the fact that it's +-1 inside a simple closed curve.\<close>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4625
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4626
lemma winding_number_homotopic_paths:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4627
    assumes "homotopic_paths (-{z}) g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4628
      shows "winding_number g z = winding_number h z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4629
proof -
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4630
  have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4631
  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4632
    using homotopic_paths_imp_subset [OF assms] by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4633
  ultimately obtain d e where "d > 0" "e > 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4634
      and d: "\<And>p. \<lbrakk>path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4635
            \<Longrightarrow> homotopic_paths (-{z}) g p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4636
      and e: "\<And>q. \<lbrakk>path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4637
            \<Longrightarrow> homotopic_paths (-{z}) h q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4638
    using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4639
  obtain p where p:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4640
       "valid_path p" "z \<notin> path_image p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4641
       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4642
       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4643
       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4644
    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4645
  obtain q where q:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4646
       "valid_path q" "z \<notin> path_image q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4647
       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4648
       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4649
       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4650
    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4651
  have gp: "homotopic_paths (- {z}) g p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4652
    by (simp add: d p valid_path_imp_path norm_minus_commute gp_less)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4653
  have hq: "homotopic_paths (- {z}) h q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4654
    by (simp add: e q valid_path_imp_path norm_minus_commute hq_less)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4655
  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4656
    apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4657
    apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4658
    apply (auto intro!: holomorphic_intros simp: p q)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4659
    done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4660
  then show ?thesis
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4661
    by (simp add: pap paq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4662
qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4663
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4664
lemma winding_number_homotopic_loops:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4665
    assumes "homotopic_loops (-{z}) g h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4666
      shows "winding_number g z = winding_number h z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4667
proof -
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4668
  have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4669
  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4670
    using homotopic_loops_imp_subset [OF assms] by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4671
  moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4672
    using homotopic_loops_imp_loop [OF assms] by auto
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4673
  ultimately obtain d e where "d > 0" "e > 0"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4674
      and d: "\<And>p. \<lbrakk>path p; pathfinish p = pathstart p; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4675
            \<Longrightarrow> homotopic_loops (-{z}) g p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4676
      and e: "\<And>q. \<lbrakk>path q; pathfinish q = pathstart q; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4677
            \<Longrightarrow> homotopic_loops (-{z}) h q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4678
    using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4679
  obtain p where p:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4680
       "valid_path p" "z \<notin> path_image p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4681
       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4682
       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4683
       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4684
    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4685
  obtain q where q:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4686
       "valid_path q" "z \<notin> path_image q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4687
       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4688
       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4689
       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4690
    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4691
  have gp: "homotopic_loops (- {z}) g p"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4692
    by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4693
  have hq: "homotopic_loops (- {z}) h q"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4694
    by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4695
  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4696
    apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"])
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4697
    apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4698
    apply (auto intro!: holomorphic_intros simp: p q)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4699
    done
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4700
  then show ?thesis
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4701
    by (simp add: pap paq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4702
qed
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4703
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4704
lemma winding_number_paths_linear_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4705
  "\<lbrakk>path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4706
    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4707
        \<Longrightarrow> winding_number h z = winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4708
  by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: )
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4709
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4710
lemma winding_number_loops_linear_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4711
  "\<lbrakk>path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4712
    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4713
        \<Longrightarrow> winding_number h z = winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4714
  by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: )
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4715
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4716
lemma winding_number_nearby_paths_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4717
     "\<lbrakk>path g; path h;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4718
      pathstart h = pathstart g; pathfinish h = pathfinish g;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4719
      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4720
      \<Longrightarrow> winding_number h z = winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4721
  by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4722
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4723
lemma winding_number_nearby_loops_eq:
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4724
     "\<lbrakk>path g; path h;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4725
      pathfinish g = pathstart g;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4726
        pathfinish h = pathstart h;
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4727
      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4728
      \<Longrightarrow> winding_number h z = winding_number g z"
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4729
  by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq)
c4f6031f1310 New material about paths, winding numbers, etc. Added lemmas to divide_const_simps. Misc tuning.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  4730
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4731
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4732
proposition winding_number_subpath_combine:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4733
    "\<lbrakk>path g; z \<notin> path_image g;
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4734
      u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4735
      \<Longrightarrow> winding_number (subpath u v g) z + winding_number (subpath v w g) z =
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4736
          winding_number (subpath u w g) z"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4737
apply (rule trans [OF winding_number_join [THEN sym]
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4738
                      winding_number_homotopic_paths [OF homotopic_join_subpaths]])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4739
apply (auto dest: path_image_subpath_subset)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4740
done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4741
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4742
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4743
subsection\<open>Partial circle path\<close>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4744
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4745
definition part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4746
  where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (ii * of_real (linepath s t x))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4747
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4748
lemma pathstart_part_circlepath [simp]:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4749
     "pathstart(part_circlepath z r s t) = z + r*exp(ii * s)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4750
by (metis part_circlepath_def pathstart_def pathstart_linepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4751
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4752
lemma pathfinish_part_circlepath [simp]:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4753
     "pathfinish(part_circlepath z r s t) = z + r*exp(ii*t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4754
by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4755
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4756
proposition has_vector_derivative_part_circlepath [derivative_intros]:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4757
    "((part_circlepath z r s t) has_vector_derivative
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4758
      (ii * r * (of_real t - of_real s) * exp(ii * linepath s t x)))
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4759
     (at x within X)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4760
  apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4761
  apply (rule has_vector_derivative_real_complex)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4762
  apply (rule derivative_eq_intros | simp)+
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4763
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4764
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4765
corollary vector_derivative_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4766
    "vector_derivative (part_circlepath z r s t) (at x) =
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4767
       ii * r * (of_real t - of_real s) * exp(ii * linepath s t x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4768
  using has_vector_derivative_part_circlepath vector_derivative_at by blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4769
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4770
corollary vector_derivative_part_circlepath01:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4771
    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4772
     \<Longrightarrow> vector_derivative (part_circlepath z r s t) (at x within {0..1}) =
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4773
          ii * r * (of_real t - of_real s) * exp(ii * linepath s t x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4774
  using has_vector_derivative_part_circlepath
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4775
  by (auto simp: vector_derivative_at_within_ivl)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4776
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4777
lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4778
  apply (simp add: valid_path_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4779
  apply (rule C1_differentiable_imp_piecewise)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4780
  apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4781
              intro!: continuous_intros)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4782
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4783
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4784
lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4785
  by (simp add: valid_path_imp_path)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4786
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4787
proposition path_image_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4788
  assumes "s \<le> t"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4789
    shows "path_image (part_circlepath z r s t) = {z + r * exp(ii * of_real x) | x. s \<le> x \<and> x \<le> t}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4790
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4791
  { fix z::real
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4792
    assume "0 \<le> z" "z \<le> 1"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4793
    with \<open>s \<le> t\<close> have "\<exists>x. (exp (\<i> * linepath s t z) = exp (\<i> * of_real x)) \<and> s \<le> x \<and> x \<le> t"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4794
      apply (rule_tac x="(1 - z) * s + z * t" in exI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4795
      apply (simp add: linepath_def scaleR_conv_of_real algebra_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4796
      apply (rule conjI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4797
      using mult_right_mono apply blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4798
      using affine_ineq  by (metis "mult.commute")
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4799
  }
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4800
  moreover
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4801
  { fix z
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4802
    assume "s \<le> z" "z \<le> t"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4803
    then have "z + of_real r * exp (\<i> * of_real z) \<in> (\<lambda>x. z + of_real r * exp (\<i> * linepath s t x)) ` {0..1}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4804
      apply (rule_tac x="(z - s)/(t - s)" in image_eqI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4805
      apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4806
      apply (auto simp: algebra_simps divide_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4807
      done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4808
  }
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4809
  ultimately show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4810
    by (fastforce simp add: path_image_def part_circlepath_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4811
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4812
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4813
corollary path_image_part_circlepath_subset:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4814
    "\<lbrakk>s \<le> t; 0 \<le> r\<rbrakk> \<Longrightarrow> path_image(part_circlepath z r s t) \<subseteq> sphere z r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4815
by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4816
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4817
proposition in_path_image_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4818
  assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4819
    shows "norm(w - z) = r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4820
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4821
  have "w \<in> {c. dist z c = r}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4822
    by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4823
  thus ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4824
    by (simp add: dist_norm norm_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4825
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4826
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4827
proposition finite_bounded_log: "finite {z::complex. norm z \<le> b \<and> exp z = w}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4828
proof (cases "w = 0")
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4829
  case True then show ?thesis by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4830
next
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4831
  case False
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4832
  have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \<i>) \<le> b + cmod (Ln w)}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4833
    apply (simp add: norm_mult finite_int_iff_bounded_le)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4834
    apply (rule_tac x="floor ((b + cmod (Ln w)) / (2*pi))" in exI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4835
    apply (auto simp: divide_simps le_floor_iff)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4836
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4837
  have [simp]: "\<And>P f. {z. P z \<and> (\<exists>n. z = f n)} = f ` {n. P (f n)}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4838
    by blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4839
  show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4840
    apply (subst exp_Ln [OF False, symmetric])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4841
    apply (simp add: exp_eq)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4842
    using norm_add_leD apply (fastforce intro: finite_subset [OF _ *])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4843
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4844
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4845
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4846
lemma finite_bounded_log2:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4847
  fixes a::complex
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4848
    assumes "a \<noteq> 0"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4849
    shows "finite {z. norm z \<le> b \<and> exp(a*z) = w}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4850
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4851
  have *: "finite ((\<lambda>z. z / a) ` {z. cmod z \<le> b * cmod a \<and> exp z = w})"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4852
    by (rule finite_imageI [OF finite_bounded_log])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4853
  show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4854
    by (rule finite_subset [OF _ *]) (force simp: assms norm_mult)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4855
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4856
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4857
proposition has_contour_integral_bound_part_circlepath_strong:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4858
  assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4859
      and "finite k" and le: "0 \<le> B" "0 < r" "s \<le> t"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4860
      and B: "\<And>x. x \<in> path_image(part_circlepath z r s t) - k \<Longrightarrow> norm(f x) \<le> B"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4861
    shows "cmod i \<le> B * r * (t - s)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4862
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4863
  consider "s = t" | "s < t" using \<open>s \<le> t\<close> by linarith
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4864
  then show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4865
  proof cases
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4866
    case 1 with fi [unfolded has_contour_integral]
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4867
    have "i = 0"  by (simp add: vector_derivative_part_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4868
    with assms show ?thesis by simp
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4869
  next
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4870
    case 2
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4871
    have [simp]: "abs r = r" using \<open>r > 0\<close> by linarith
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4872
    have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4873
      by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4874
    have "finite (part_circlepath z r s t -` {y} \<inter> {0..1})" if "y \<in> k" for y
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4875
    proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4876
      def w \<equiv> "(y - z)/of_real r / exp(ii * of_real s)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4877
      have fin: "finite (of_real -` {z. cmod z \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4878
        apply (rule finite_vimageI [OF finite_bounded_log2])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4879
        using \<open>s < t\<close> apply (auto simp: inj_of_real)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4880
        done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4881
      show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4882
        apply (simp add: part_circlepath_def linepath_def vimage_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4883
        apply (rule finite_subset [OF _ fin])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4884
        using le
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4885
        apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4886
        done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4887
    qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4888
    then have fin01: "finite ((part_circlepath z r s t) -` k \<inter> {0..1})"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4889
      by (rule finite_finite_vimage_IntI [OF \<open>finite k\<close>])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4890
    have **: "((\<lambda>x. if (part_circlepath z r s t x) \<in> k then 0
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4891
                    else f(part_circlepath z r s t x) *
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4892
                       vector_derivative (part_circlepath z r s t) (at x)) has_integral i)  {0..1}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4893
      apply (rule has_integral_spike
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4894
              [where f = "\<lambda>x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4895
      apply (rule negligible_finite [OF fin01])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4896
      using fi has_contour_integral
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4897
      apply auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4898
      done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4899
    have *: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1; part_circlepath z r s t x \<notin> k\<rbrakk> \<Longrightarrow> cmod (f (part_circlepath z r s t x)) \<le> B"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4900
      by (auto intro!: B [unfolded path_image_def image_def, simplified])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4901
    show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4902
      apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4903
      using assms apply force
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4904
      apply (simp add: norm_mult vector_derivative_part_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4905
      using le * "2" \<open>r > 0\<close> by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4906
  qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4907
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4908
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4909
corollary has_contour_integral_bound_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4910
      "\<lbrakk>(f has_contour_integral i) (part_circlepath z r s t);
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4911
        0 \<le> B; 0 < r; s \<le> t;
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4912
        \<And>x. x \<in> path_image(part_circlepath z r s t) \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4913
       \<Longrightarrow> norm i \<le> B*r*(t - s)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4914
  by (auto intro: has_contour_integral_bound_part_circlepath_strong)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4915
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4916
proposition contour_integrable_continuous_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4917
     "continuous_on (path_image (part_circlepath z r s t)) f
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4918
      \<Longrightarrow> f contour_integrable_on (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4919
  apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4920
  apply (rule integrable_continuous_real)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4921
  apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4922
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4923
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4924
proposition winding_number_part_circlepath_pos_less:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4925
  assumes "s < t" and no: "norm(w - z) < r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4926
    shows "0 < Re (winding_number(part_circlepath z r s t) w)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4927
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4928
  have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4929
  note valid_path_part_circlepath
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4930
  moreover have " w \<notin> path_image (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4931
    using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4932
  moreover have "0 < r * (t - s) * (r - cmod (w - z))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4933
    using assms by (metis \<open>0 < r\<close> diff_gt_0_iff_gt mult_pos_pos)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4934
  ultimately show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4935
    apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4936
    apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4937
    apply (rule mult_left_mono)+
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4938
    using Re_Im_le_cmod [of "w-z" "linepath s t x" for x]
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4939
    apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4940
    using assms \<open>0 < r\<close> by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4941
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4942
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4943
proposition simple_path_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4944
    "simple_path(part_circlepath z r s t) \<longleftrightarrow> (r \<noteq> 0 \<and> s \<noteq> t \<and> abs(s - t) \<le> 2*pi)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4945
proof (cases "r = 0 \<or> s = t")
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4946
  case True
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4947
  then show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4948
    apply (rule disjE)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4949
    apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4950
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4951
next
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4952
  case False then have "r \<noteq> 0" "s \<noteq> t" by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4953
  have *: "\<And>x y z s t. ii*((1 - x) * s + x * t) = ii*(((1 - y) * s + y * t)) + z  \<longleftrightarrow> ii*(x - y) * (t - s) = z"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4954
    by (simp add: algebra_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4955
  have abs01: "\<And>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4956
                      \<Longrightarrow> (x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0 \<longleftrightarrow> abs(x - y) \<in> {0,1})"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4957
    by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4958
  have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P(abs(x - y))) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4959
    by force
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4960
  have **: "\<And>x y. (\<exists>n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \<longleftrightarrow>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4961
                  (\<exists>n. abs(x - y) * (t - s) = 2 * (of_int n * pi))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4962
    by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real]
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4963
                    intro: exI [where x = "-n" for n])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4964
  have 1: "\<forall>x. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> (\<exists>n. x * (t - s) = 2 * (real_of_int n * pi)) \<longrightarrow> x = 0 \<or> x = 1 \<Longrightarrow> \<bar>s - t\<bar> \<le> 2 * pi"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4965
    apply (rule ccontr)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4966
    apply (drule_tac x="2*pi / abs(t-s)" in spec)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4967
    using False
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4968
    apply (simp add: abs_minus_commute divide_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4969
    apply (frule_tac x=1 in spec)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4970
    apply (drule_tac x="-1" in spec)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4971
    apply (simp add:)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4972
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4973
  have 2: "\<bar>s - t\<bar> = \<bar>2 * (real_of_int n * pi) / x\<bar>" if "x \<noteq> 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4974
  proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4975
    have "t-s = 2 * (real_of_int n * pi)/x"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4976
      using that by (simp add: field_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4977
    then show ?thesis by (metis abs_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4978
  qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4979
  show ?thesis using False
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4980
    apply (simp add: simple_path_def path_part_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4981
    apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01  del: Set.insert_iff)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4982
    apply (subst abs_away)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4983
    apply (auto simp: 1)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4984
    apply (rule ccontr)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4985
    apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4986
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4987
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4988
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4989
proposition arc_part_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4990
  assumes "r \<noteq> 0" "s \<noteq> t" "abs(s - t) < 2*pi"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4991
    shows "arc (part_circlepath z r s t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4992
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4993
  have *: "x = y" if eq: "\<i> * (linepath s t x) = \<i> * (linepath s t y) + 2 * of_int n * complex_of_real pi * \<i>"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4994
                  and x: "x \<in> {0..1}" and y: "y \<in> {0..1}" for x y n
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4995
    proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4996
      have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4997
        by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4998
      then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  4999
        by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5000
      then have st: "x \<noteq> y \<Longrightarrow> (s-t) = (of_int n * (pi * 2) / (y-x))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5001
        by (force simp: field_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5002
      show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5003
        apply (rule ccontr)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5004
        using assms x y
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5005
        apply (simp add: st abs_mult field_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5006
        using st
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5007
        apply (auto simp: dest: of_int_lessD)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5008
        done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5009
    qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5010
  show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5011
    using assms
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5012
    apply (simp add: arc_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5013
    apply (simp add: part_circlepath_def inj_on_def exp_eq)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5014
    apply (blast intro: *)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5015
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5016
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5017
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5018
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5019
subsection\<open>Special case of one complete circle\<close>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5020
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5021
definition circlepath :: "[complex, real, real] \<Rightarrow> complex"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5022
  where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5023
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5024
lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * ii * of_real x))"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5025
  by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5026
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5027
lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5028
  by (simp add: circlepath_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5029
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5030
lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5031
  by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5032
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5033
proposition has_vector_derivative_circlepath [derivative_intros]:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5034
 "((circlepath z r) has_vector_derivative (2 * pi * ii * r * exp (2 * of_real pi * ii * of_real x)))
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5035
   (at x within X)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5036
  apply (simp add: circlepath_def scaleR_conv_of_real)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5037
  apply (rule derivative_eq_intros)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5038
  apply (simp add: algebra_simps)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5039
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5040
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5041
corollary vector_derivative_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5042
   "vector_derivative (circlepath z r) (at x) =
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5043
    2 * pi * ii * r * exp(2 * of_real pi * ii * x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5044
using has_vector_derivative_circlepath vector_derivative_at by blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5045
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5046
corollary vector_derivative_circlepath01:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5047
    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5048
     \<Longrightarrow> vector_derivative (circlepath z r) (at x within {0..1}) =
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5049
          2 * pi * ii * r * exp(2 * of_real pi * ii * x)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5050
  using has_vector_derivative_circlepath
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5051
  by (auto simp: vector_derivative_at_within_ivl)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5052
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5053
lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5054
  by (simp add: circlepath_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5055
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5056
lemma path_circlepath [simp]: "path (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5057
  by (simp add: valid_path_imp_path)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5058
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5059
proposition path_image_circlepath [simp]:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5060
  assumes "0 \<le> r" shows "path_image (circlepath z r) = sphere z r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5061
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5062
  have *: "x \<in> (\<lambda>u. z + (cmod (x - z)) * exp (\<i> * (of_real u * (of_real pi * 2)))) ` {0..1}" for x
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5063
  proof (cases "x = z")
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5064
    case True then show ?thesis by force
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5065
  next
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5066
    case False
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5067
    def w \<equiv> "x - z"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5068
    then have "w \<noteq> 0" by (simp add: False)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5069
    have **: "\<And>t. \<lbrakk>Re w = cos t * cmod w; Im w = sin t * cmod w\<rbrakk> \<Longrightarrow> w = of_real (cmod w) * exp (\<i> * t)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5070
      using cis_conv_exp complex_eq_iff by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5071
    show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5072
      apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5073
      apply (simp add: divide_simps \<open>w \<noteq> 0\<close> cmod_power2 [symmetric])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5074
      apply (rule_tac x="t / (2*pi)" in image_eqI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5075
      apply (simp add: divide_simps \<open>w \<noteq> 0\<close>)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5076
      using False **
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5077
      apply (auto simp: w_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5078
      done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5079
  qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5080
  show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5081
    unfolding circlepath path_image_def sphere_def dist_norm
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5082
    by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5083
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5084
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5085
lemma has_contour_integral_bound_circlepath_strong:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5086
      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5087
        finite k; 0 \<le> B; 0 < r;
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5088
        \<And>x. \<lbrakk>norm(x - z) = r; x \<notin> k\<rbrakk> \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5089
        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5090
  unfolding circlepath_def
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5091
  by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5092
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5093
corollary has_contour_integral_bound_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5094
      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5095
        0 \<le> B; 0 < r; \<And>x. norm(x - z) = r \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5096
        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5097
  by (auto intro: has_contour_integral_bound_circlepath_strong)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5098
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5099
proposition contour_integrable_continuous_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5100
    "continuous_on (path_image (circlepath z r)) f
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5101
     \<Longrightarrow> f contour_integrable_on (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5102
  by (simp add: circlepath_def contour_integrable_continuous_part_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5103
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5104
lemma simple_path_circlepath: "simple_path(circlepath z r) \<longleftrightarrow> (r \<noteq> 0)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5105
  by (simp add: circlepath_def simple_path_part_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5106
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5107
lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \<Longrightarrow> w \<notin> path_image (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5108
  apply (subst path_image_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5109
  apply (meson le_cases less_le_trans norm_not_less_zero)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5110
  apply (simp add: sphere_def dist_norm norm_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5111
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5112
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5113
proposition contour_integral_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5114
     "0 < r \<Longrightarrow> contour_integral (circlepath z r) (\<lambda>w. 1 / (w - z)) = 2 * complex_of_real pi * \<i>"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5115
  apply (rule contour_integral_unique)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5116
  apply (simp add: has_contour_integral_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5117
  apply (subst has_integral_cong)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5118
  apply (simp add: vector_derivative_circlepath01)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5119
  using has_integral_const_real [of _ 0 1]
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5120
  apply (force simp: circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5121
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5122
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5123
lemma winding_number_circlepath_centre: "0 < r \<Longrightarrow> winding_number (circlepath z r) z = 1"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5124
  apply (rule winding_number_unique_loop)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5125
  apply (simp_all add: sphere_def valid_path_imp_path)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5126
  apply (rule_tac x="circlepath z r" in exI)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5127
  apply (simp add: sphere_def contour_integral_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5128
  done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5129
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5130
proposition winding_number_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5131
  assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5132
proof (cases "w = z")
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5133
  case True then show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5134
    using assms winding_number_circlepath_centre by auto
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5135
next
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5136
  case False
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5137
  have [simp]: "r > 0"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5138
    using assms le_less_trans norm_ge_zero by blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5139
  def r' \<equiv> "norm(w - z)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5140
  have "r' < r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5141
    by (simp add: assms r'_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5142
  have disjo: "cball z r' \<inter> sphere z r = {}"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5143
    using \<open>r' < r\<close> by (force simp: cball_def sphere_def)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5144
  have "winding_number(circlepath z r) w = winding_number(circlepath z r) z"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5145
    apply (rule winding_number_around_inside [where s = "cball z r'"])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5146
    apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5147
    apply (simp_all add: False r'_def dist_norm norm_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5148
    done
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5149
  also have "... = 1"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5150
    by (simp add: winding_number_circlepath_centre)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5151
  finally show ?thesis .
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5152
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5153
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5154
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5155
text\<open> Hence the Cauchy formula for points inside a circle.\<close>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5156
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5157
theorem Cauchy_integral_circlepath:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5158
  assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5159
  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * ii * f w))
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5160
         (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5161
proof -
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5162
  have "r > 0"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5163
    using assms le_less_trans norm_ge_zero by blast
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5164
  have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5165
        (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5166
    apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5167
    using assms  \<open>r > 0\<close>
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5168
    apply (simp_all add: dist_norm norm_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5169
    apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5170
    apply (simp add: cball_def sphere_def dist_norm, clarify)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5171
    apply (simp add:)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5172
    by (metis dist_commute dist_norm less_irrefl)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5173
  then show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5174
    by (simp add: winding_number_circlepath assms)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5175
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5176
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5177
corollary Cauchy_integral_circlepath_simple:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5178
  assumes "f holomorphic_on cball z r" "norm(w - z) < r"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5179
  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * ii * f w))
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5180
         (circlepath z r)"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5181
using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5182
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5183
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5184
lemma no_bounded_connected_component_imp_winding_number_zero:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5185
  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5186
      and nb: "\<And>z. bounded (connected_component_set (- s) z) \<longrightarrow> z \<in> s"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5187
  shows "winding_number g z = 0"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5188
apply (rule winding_number_zero_in_outside)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5189
apply (simp_all add: assms)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5190
by (metis nb [of z] \<open>path_image g \<subseteq> s\<close> \<open>z \<notin> s\<close> contra_subsetD mem_Collect_eq outside outside_mono)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5191
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5192
lemma no_bounded_path_component_imp_winding_number_zero:
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5193
  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5194
      and nb: "\<And>z. bounded (path_component_set (- s) z) \<longrightarrow> z \<in> s"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5195
  shows "winding_number g z = 0"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5196
apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g])
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5197
by (simp add: bounded_subset nb path_component_subset_connected_component)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  5198
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  5199
end