merged
authorwenzelm
Fri, 04 Sep 2015 22:16:54 +0200
changeset 61117 4b5872b9d783
parent 61108 279a5b4f47bd (diff)
parent 61116 6189d179c2b5 (current diff)
child 61118 a39e9c46a772
merged
--- a/src/HOL/Complex.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Complex.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -527,6 +527,9 @@
    (auto simp: complex_eq_iff norm_complex_def power2_eq_square[symmetric] of_real_power[symmetric]
          simp del: of_real_power)
 
+lemma complex_div_cnj: "a / b = (a * cnj b) / (norm b)^2"
+  using complex_norm_square by auto
+
 lemma Re_complex_div_eq_0: "Re (a / b) = 0 \<longleftrightarrow> Re (a * cnj b) = 0"
   by (auto simp add: Re_divide)
 
@@ -567,6 +570,18 @@
 lemma Im_complex_div_le_0: "Im(a / b) \<le> 0 \<longleftrightarrow> Im(a * cnj b) \<le> 0"
   by (metis Im_complex_div_gt_0 not_le)
 
+lemma Re_divide_of_real [simp]: "Re (z / of_real r) = Re z / r"
+  by (simp add: Re_divide power2_eq_square)
+
+lemma Im_divide_of_real [simp]: "Im (z / of_real r) = Im z / r"
+  by (simp add: Im_divide power2_eq_square)
+
+lemma Re_divide_Reals: "r \<in> Reals \<Longrightarrow> Re (z / r) = Re z / Re r"
+  by (metis Re_divide_of_real of_real_Re)
+
+lemma Im_divide_Reals: "r \<in> Reals \<Longrightarrow> Im (z / r) = Im z / Re r"
+  by (metis Im_divide_of_real of_real_Re)
+
 lemma Re_setsum[simp]: "Re (setsum f s) = (\<Sum>x\<in>s. Re (f x))"
   by (induct s rule: infinite_finite_induct) auto
 
@@ -588,6 +603,12 @@
 lemma summable_Im: "summable f \<Longrightarrow> summable (\<lambda>x. Im (f x))"
   unfolding summable_complex_iff by blast
 
+lemma complex_is_Nat_iff: "z \<in> \<nat> \<longleftrightarrow> Im z = 0 \<and> (\<exists>i. Re z = of_nat i)"
+  by (auto simp: Nats_def complex_eq_iff)
+
+lemma complex_is_Int_iff: "z \<in> \<int> \<longleftrightarrow> Im z = 0 \<and> (\<exists>i. Re z = of_int i)"
+  by (auto simp: Ints_def complex_eq_iff)
+
 lemma complex_is_Real_iff: "z \<in> \<real> \<longleftrightarrow> Im z = 0"
   by (auto simp: Reals_def complex_eq_iff)
 
--- a/src/HOL/Library/Fraction_Field.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Library/Fraction_Field.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -13,24 +13,23 @@
 
 subsubsection \<open>Construction of the type of fractions\<close>
 
-definition fractrel :: "(('a::idom * 'a ) * ('a * 'a)) set" where
-  "fractrel = {(x, y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
+context idom begin
+
+definition fractrel :: "'a \<times> 'a \<Rightarrow> 'a * 'a \<Rightarrow> bool" where
+  "fractrel = (\<lambda>x y. snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x)"
 
 lemma fractrel_iff [simp]:
-  "(x, y) \<in> fractrel \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
+  "fractrel x y \<longleftrightarrow> snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x"
   by (simp add: fractrel_def)
 
-lemma refl_fractrel: "refl_on {x. snd x \<noteq> 0} fractrel"
-  by (auto simp add: refl_on_def fractrel_def)
-
-lemma sym_fractrel: "sym fractrel"
-  by (simp add: fractrel_def sym_def)
+lemma symp_fractrel: "symp fractrel"
+  by (simp add: symp_def)
 
-lemma trans_fractrel: "trans fractrel"
-proof (rule transI, unfold split_paired_all)
+lemma transp_fractrel: "transp fractrel"
+proof (rule transpI, unfold split_paired_all)
   fix a b a' b' a'' b'' :: 'a
-  assume A: "((a, b), (a', b')) \<in> fractrel"
-  assume B: "((a', b'), (a'', b'')) \<in> fractrel"
+  assume A: "fractrel (a, b) (a', b')"
+  assume B: "fractrel (a', b') (a'', b'')"
   have "b' * (a * b'') = b'' * (a * b')" by (simp add: ac_simps)
   also from A have "a * b' = a' * b" by auto
   also have "b'' * (a' * b) = b * (a' * b'')" by (simp add: ac_simps)
@@ -39,46 +38,27 @@
   finally have "b' * (a * b'') = b' * (a'' * b)" .
   moreover from B have "b' \<noteq> 0" by auto
   ultimately have "a * b'' = a'' * b" by simp
-  with A B show "((a, b), (a'', b'')) \<in> fractrel" by auto
+  with A B show "fractrel (a, b) (a'', b'')" by auto
 qed
 
-lemma equiv_fractrel: "equiv {x. snd x \<noteq> 0} fractrel"
-  by (rule equivI [OF refl_fractrel sym_fractrel trans_fractrel])
-
-lemmas UN_fractrel = UN_equiv_class [OF equiv_fractrel]
-lemmas UN_fractrel2 = UN_equiv_class2 [OF equiv_fractrel equiv_fractrel]
-
-lemma equiv_fractrel_iff [iff]:
-  assumes "snd x \<noteq> 0" and "snd y \<noteq> 0"
-  shows "fractrel `` {x} = fractrel `` {y} \<longleftrightarrow> (x, y) \<in> fractrel"
-  by (rule eq_equiv_class_iff, rule equiv_fractrel) (auto simp add: assms)
-
-definition "fract = {(x::'a\<times>'a). snd x \<noteq> (0::'a::idom)} // fractrel"
+lemma part_equivp_fractrel: "part_equivp fractrel"
+using _ symp_fractrel transp_fractrel
+by(rule part_equivpI)(rule exI[where x="(0, 1)"]; simp)
 
-typedef 'a fract = "fract :: ('a * 'a::idom) set set"
-  unfolding fract_def
-proof
-  have "(0::'a, 1::'a) \<in> {x. snd x \<noteq> 0}" by simp
-  then show "fractrel `` {(0::'a, 1)} \<in> {x. snd x \<noteq> 0} // fractrel"
-    by (rule quotientI)
-qed
+end
 
-lemma fractrel_in_fract [simp]: "snd x \<noteq> 0 \<Longrightarrow> fractrel `` {x} \<in> fract"
-  by (simp add: fract_def quotientI)
-
-declare Abs_fract_inject [simp] Abs_fract_inverse [simp]
-
+quotient_type 'a fract = "'a :: idom \<times> 'a" / partial: "fractrel"
+by(rule part_equivp_fractrel)
 
 subsubsection \<open>Representation and basic operations\<close>
 
-definition Fract :: "'a::idom \<Rightarrow> 'a \<Rightarrow> 'a fract"
-  where "Fract a b = Abs_fract (fractrel `` {if b = 0 then (0, 1) else (a, b)})"
-
-code_datatype Fract
+lift_definition Fract :: "'a :: idom \<Rightarrow> 'a \<Rightarrow> 'a fract"
+  is "\<lambda>a b. if b = 0 then (0, 1) else (a, b)"
+  by simp
 
 lemma Fract_cases [cases type: fract]:
   obtains (Fract) a b where "q = Fract a b" "b \<noteq> 0"
-  by (cases q) (clarsimp simp add: Fract_def fract_def quotient_def)
+by transfer simp
 
 lemma Fract_induct [case_names Fract, induct type: fract]:
   "(\<And>a b. b \<noteq> 0 \<Longrightarrow> P (Fract a b)) \<Longrightarrow> P q"
@@ -88,40 +68,37 @@
   shows "\<And>a b c d. b \<noteq> 0 \<Longrightarrow> d \<noteq> 0 \<Longrightarrow> Fract a b = Fract c d \<longleftrightarrow> a * d = c * b"
     and "\<And>a. Fract a 0 = Fract 0 1"
     and "\<And>a c. Fract 0 a = Fract 0 c"
-  by (simp_all add: Fract_def)
+by(transfer; simp)+
 
 instantiation fract :: (idom) "{comm_ring_1,power}"
 begin
 
-definition Zero_fract_def [code_unfold]: "0 = Fract 0 1"
+lift_definition zero_fract :: "'a fract" is "(0, 1)" by simp
 
-definition One_fract_def [code_unfold]: "1 = Fract 1 1"
+lemma Zero_fract_def: "0 = Fract 0 1"
+by transfer simp
+
+lift_definition one_fract :: "'a fract" is "(1, 1)" by simp
 
-definition add_fract_def:
-  "q + r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
-    fractrel `` {(fst x * snd y + fst y * snd x, snd x * snd y)})"
+lemma One_fract_def: "1 = Fract 1 1"
+by transfer simp
+
+lift_definition plus_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>q r. (fst q * snd r + fst r * snd q, snd q * snd r)"
+by(auto simp add: algebra_simps)
 
 lemma add_fract [simp]:
-  assumes "b \<noteq> (0::'a::idom)"
-    and "d \<noteq> 0"
-  shows "Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
-proof -
-  have "(\<lambda>x y. fractrel``{(fst x * snd y + fst y * snd x, snd x * snd y :: 'a)}) respects2 fractrel"
-    by (rule equiv_fractrel [THEN congruent2_commuteI]) (simp_all add: algebra_simps)
-  with assms show ?thesis by (simp add: Fract_def add_fract_def UN_fractrel2)
-qed
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b + Fract c d = Fract (a * d + c * b) (b * d)"
+by transfer simp
 
-definition minus_fract_def:
-  "- q = Abs_fract (\<Union>x \<in> Rep_fract q. fractrel `` {(- fst x, snd x)})"
+lift_definition uminus_fract :: "'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>x. (- fst x, snd x)"
+by simp
 
-lemma minus_fract [simp, code]:
+lemma minus_fract [simp]:
   fixes a b :: "'a::idom"
   shows "- Fract a b = Fract (- a) b"
-proof -
-  have "(\<lambda>x. fractrel `` {(- fst x, snd x :: 'a)}) respects fractrel"
-    by (simp add: congruent_def split_paired_all)
-  then show ?thesis by (simp add: Fract_def minus_fract_def UN_fractrel)
-qed
+by transfer simp
 
 lemma minus_fract_cancel [simp]: "Fract (- a) (- b) = Fract a b"
   by (cases "b = 0") (simp_all add: eq_fract)
@@ -129,31 +106,19 @@
 definition diff_fract_def: "q - r = q + - (r::'a fract)"
 
 lemma diff_fract [simp]:
-  assumes "b \<noteq> 0"
-    and "d \<noteq> 0"
-  shows "Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
-  using assms by (simp add: diff_fract_def)
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b - Fract c d = Fract (a * d - c * b) (b * d)"
+  by (simp add: diff_fract_def)
 
-definition mult_fract_def:
-  "q * r = Abs_fract (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
-    fractrel``{(fst x * fst y, snd x * snd y)})"
+lift_definition times_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>q r. (fst q * fst r, snd q * snd r)"
+by(simp add: algebra_simps)
 
 lemma mult_fract [simp]: "Fract (a::'a::idom) b * Fract c d = Fract (a * c) (b * d)"
-proof -
-  have "(\<lambda>x y. fractrel `` {(fst x * fst y, snd x * snd y :: 'a)}) respects2 fractrel"
-    by (rule equiv_fractrel [THEN congruent2_commuteI]) (simp_all add: algebra_simps)
-  then show ?thesis by (simp add: Fract_def mult_fract_def UN_fractrel2)
-qed
+by transfer simp
 
 lemma mult_fract_cancel:
-  assumes "c \<noteq> (0::'a)"
-  shows "Fract (c * a) (c * b) = Fract a b"
-proof -
-  from assms have "Fract c c = Fract 1 1"
-    by (simp add: Fract_def)
-  then show ?thesis
-    by (simp add: mult_fract [symmetric])
-qed
+  "c \<noteq> 0 \<Longrightarrow> Fract (c * a) (c * b) = Fract a b"
+by transfer simp
 
 instance
 proof
@@ -188,14 +153,13 @@
 lemma Fract_of_nat_eq: "Fract (of_nat k) 1 = of_nat k"
   by (rule of_nat_fract [symmetric])
 
-lemma fract_collapse [code_post]:
+lemma fract_collapse:
   "Fract 0 k = 0"
   "Fract 1 1 = 1"
   "Fract k 0 = 0"
-  by (cases "k = 0")
-    (simp_all add: Zero_fract_def One_fract_def eq_fract Fract_def)
+by(transfer; simp)+
 
-lemma fract_expand [code_unfold]:
+lemma fract_expand:
   "0 = Fract 0 1"
   "1 = Fract 1 1"
   by (simp_all add: fract_collapse)
@@ -227,19 +191,12 @@
 instantiation fract :: (idom) field
 begin
 
-definition inverse_fract_def:
-  "inverse q = Abs_fract (\<Union>x \<in> Rep_fract q.
-     fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x)})"
+lift_definition inverse_fract :: "'a fract \<Rightarrow> 'a fract"
+  is "\<lambda>x. if fst x = 0 then (0, 1) else (snd x, fst x)"
+by(auto simp add: algebra_simps)
 
 lemma inverse_fract [simp]: "inverse (Fract a b) = Fract (b::'a::idom) a"
-proof -
-  have *: "\<And>x. (0::'a) = x \<longleftrightarrow> x = 0"
-    by auto
-  have "(\<lambda>x. fractrel `` {if fst x = 0 then (0, 1) else (snd x, fst x :: 'a)}) respects fractrel"
-    by (auto simp add: congruent_def * algebra_simps)
-  then show ?thesis
-    by (simp add: Fract_def inverse_fract_def UN_fractrel)
-qed
+by transfer simp
 
 definition divide_fract_def: "q div r = q * inverse (r:: 'a fract)"
 
@@ -266,16 +223,16 @@
 
 subsubsection \<open>The ordered field of fractions over an ordered idom\<close>
 
-lemma le_congruent2:
-  "(\<lambda>x y::'a \<times> 'a::linordered_idom.
-    {(fst x * snd y)*(snd x * snd y) \<le> (fst y * snd x)*(snd x * snd y)})
-    respects2 fractrel"
-proof (clarsimp simp add: congruent2_def)
-  fix a b a' b' c d c' d' :: 'a
-  assume neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
-  assume eq1: "a * b' = a' * b"
-  assume eq2: "c * d' = c' * d"
+instantiation fract :: (linordered_idom) linorder
+begin
 
+lemma less_eq_fract_respect:
+  fixes a b a' b' c d c' d' :: 'a
+  assumes neq: "b \<noteq> 0"  "b' \<noteq> 0"  "d \<noteq> 0"  "d' \<noteq> 0"
+  assumes eq1: "a * b' = a' * b"
+  assumes eq2: "c * d' = c' * d"
+  shows "((a * d) * (b * d) \<le> (c * b) * (b * d)) \<longleftrightarrow> ((a' * d') * (b' * d') \<le> (c' * b') * (b' * d'))"
+proof -
   let ?le = "\<lambda>a b c d. ((a * d) * (b * d) \<le> (c * b) * (b * d))"
   {
     fix a b c d x :: 'a
@@ -309,25 +266,18 @@
   finally show "?le a b c d = ?le a' b' c' d'" .
 qed
 
-instantiation fract :: (linordered_idom) linorder
-begin
-
-definition le_fract_def:
-  "q \<le> r \<longleftrightarrow> the_elem (\<Union>x \<in> Rep_fract q. \<Union>y \<in> Rep_fract r.
-    {(fst x * snd y) * (snd x * snd y) \<le> (fst y * snd x) * (snd x * snd y)})"
+lift_definition less_eq_fract :: "'a fract \<Rightarrow> 'a fract \<Rightarrow> bool"
+  is "\<lambda>q r. (fst q * snd r) * (snd q * snd r) \<le> (fst r * snd q) * (snd q * snd r)"
+by (clarsimp simp add: less_eq_fract_respect)
 
 definition less_fract_def: "z < (w::'a fract) \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
 
 lemma le_fract [simp]:
-  assumes "b \<noteq> 0"
-    and "d \<noteq> 0"
-  shows "Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
-  by (simp add: Fract_def le_fract_def le_congruent2 UN_fractrel2 assms)
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b \<le> Fract c d \<longleftrightarrow> (a * d) * (b * d) \<le> (c * b) * (b * d)"
+  by transfer simp
 
 lemma less_fract [simp]:
-  assumes "b \<noteq> 0"
-    and "d \<noteq> 0"
-  shows "Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
+  "\<lbrakk> b \<noteq> 0; d \<noteq> 0 \<rbrakk> \<Longrightarrow> Fract a b < Fract c d \<longleftrightarrow> (a * d) * (b * d) < (c * b) * (b * d)"
   by (simp add: less_fract_def less_le_not_le ac_simps assms)
 
 instance
@@ -406,14 +356,14 @@
 instantiation fract :: (linordered_idom) "{distrib_lattice,abs_if,sgn_if}"
 begin
 
-definition abs_fract_def: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
+definition abs_fract_def2: "\<bar>q\<bar> = (if q < 0 then -q else (q::'a fract))"
 
 definition sgn_fract_def:
   "sgn (q::'a fract) = (if q = 0 then 0 else if 0 < q then 1 else - 1)"
 
 theorem abs_fract [simp]: "\<bar>Fract a b\<bar> = Fract \<bar>a\<bar> \<bar>b\<bar>"
-  by (auto simp add: abs_fract_def Zero_fract_def le_less
-      eq_fract zero_less_mult_iff mult_less_0_iff split: abs_split)
+unfolding abs_fract_def2 not_le[symmetric]
+by transfer(auto simp add: zero_less_mult_iff le_less)
 
 definition inf_fract_def:
   "(inf :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = min"
@@ -422,9 +372,7 @@
   "(sup :: 'a fract \<Rightarrow> 'a fract \<Rightarrow> 'a fract) = max"
 
 instance
-  by intro_classes
-    (auto simp add: abs_fract_def sgn_fract_def
-      max_min_distrib2 inf_fract_def sup_fract_def)
+by intro_classes (simp_all add: abs_fract_def2 sgn_fract_def inf_fract_def sup_fract_def max_min_distrib2)
 
 end
 
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -1140,10 +1140,6 @@
 
 subsection "Derivative"
 
-lemma differentiable_at_imp_differentiable_on:
-  "(\<forall>x\<in>(s::(real^'n) set). f differentiable at x) \<Longrightarrow> f differentiable_on s"
-  by (metis differentiable_at_withinI differentiable_on_def)
-
 definition "jacobian f net = matrix(frechet_derivative f net)"
 
 lemma jacobian_works:
--- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -1,7 +1,7 @@
 section \<open>Complex path integrals and Cauchy's integral theorem\<close>
 
 theory Cauchy_Integral_Thm
-imports Complex_Transcendental Path_Connected
+imports Complex_Transcendental Weierstrass
 begin
 
 
@@ -2512,13 +2512,14 @@
   apply (rule path_integrable_holomorphic [OF contf os Finite_Set.finite.emptyI g])
   using fh  by (simp add: complex_differentiable_def holomorphic_on_open os)
 
-lemma path_integrable_inversediff:
+lemma continuous_on_inversediff:
+  fixes z:: "'a::real_normed_field" shows "z \<notin> s \<Longrightarrow> continuous_on s (\<lambda>w. 1 / (w - z))"
+  by (rule continuous_intros | force)+
+
+corollary path_integrable_inversediff:
     "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) path_integrable_on g"
-apply (rule path_integrable_holomorphic_simple [of "UNIV-{z}"])
-    apply (rule continuous_intros | simp)+
- apply blast
-apply (simp add: holomorphic_on_open open_delete)
-apply (force intro: derivative_eq_intros)
+apply (rule path_integrable_holomorphic_simple [of "UNIV-{z}", OF continuous_on_inversediff])
+apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros)
 done
 
 text{*Key fact that path integral is the same for a "nearby" path. This is the
@@ -2688,7 +2689,7 @@
                  \<subseteq> ball (p t) (ee (p t))"
             apply (intro subset_path_image_join pi_hgn pi_ghn')
             using `N>0` Suc.prems
-            apply (auto simp: dist_norm field_simps ptgh_ee)
+            apply (auto simp: dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee)
             done
           have pi0: "(f has_path_integral 0)
                        (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++
@@ -2778,4 +2779,64 @@
   using path_integral_nearby [OF assms, where Ends=False]
   by simp_all
 
+lemma valid_path_polynomial_function:
+  fixes p :: "real \<Rightarrow> 'b::euclidean_space"
+  shows "polynomial_function p \<Longrightarrow> valid_path p"
+apply (simp add: valid_path_def)
+apply (rule differentiable_on_imp_piecewise_differentiable [OF differentiable_at_imp_differentiable_on])
+using differentiable_def has_vector_derivative_def
+apply (blast intro: dest: has_vector_derivative_polynomial_function)
+done
+
+lemma path_integral_bound_exists:
+assumes s: "open s"
+    and g: "valid_path g"
+    and pag: "path_image g \<subseteq> s"
+  shows "\<exists>L. 0 < L \<and>
+	     (\<forall>f B. f holomorphic_on s \<and> (\<forall>z \<in> s. norm(f z) \<le> B)
+		     \<longrightarrow> norm(path_integral g f) \<le> L*B)"
+proof -
+have "path g" using g
+  by (simp add: valid_path_imp_path)
+then obtain d::real and p
+  where d: "0 < d"
+    and p: "polynomial_function p" "path_image p \<subseteq> s"
+    and pi: "\<And>f. f holomorphic_on s \<Longrightarrow> path_integral g f = path_integral p f"
+  using path_integral_nearby_ends [OF s `path g` pag]
+  apply clarify
+  apply (drule_tac x=g in spec)
+  apply (simp only: assms)
+  apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function)
+  done
+then obtain p' where p': "polynomial_function p'"
+	       "\<And>x. (p has_vector_derivative (p' x)) (at x)"
+  using has_vector_derivative_polynomial_function by force
+then have "bounded(p' ` {0..1})"
+  using continuous_on_polymonial_function
+  by (force simp: intro!: compact_imp_bounded compact_continuous_image)
+then obtain L where L: "L>0" and nop': "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (p' x) \<le> L"
+  by (force simp: bounded_pos)
+{ fix f B
+  assume f: "f holomorphic_on s"
+     and B: "\<And>z. z\<in>s \<Longrightarrow> cmod (f z) \<le> B"
+  then have "f path_integrable_on p \<and> valid_path p"
+    using p s
+    by (blast intro: valid_path_polynomial_function path_integrable_holomorphic_simple holomorphic_on_imp_continuous_on)
+  moreover have "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (vector_derivative p (at x)) * cmod (f (p x)) \<le> L * B"
+    apply (rule mult_mono)
+    apply (subst Derivative.vector_derivative_at; force intro: p' nop')
+    using L B p
+    apply (auto simp: path_image_def image_subset_iff)
+    done
+  ultimately have "cmod (path_integral g f) \<le> L * B"
+    apply (simp add: pi [OF f])
+    apply (simp add: path_integral_integral)
+    apply (rule order_trans [OF integral_norm_bound_integral])
+    apply (auto simp: mult.commute integral_norm_bound_integral path_integrable_on [symmetric] norm_mult)
+    done
+} then
+show ?thesis
+  by (force simp: L path_integral_integral)
+qed
+
 end
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -6375,7 +6375,7 @@
   using segment_furthest_le[OF assms, of b]
   by (auto simp add:norm_minus_commute)
 
-lemma segment_refl: "closed_segment a a = {a}"
+lemma segment_refl [simp]: "closed_segment a a = {a}"
   unfolding segment by (auto simp add: algebra_simps)
 
 lemma closed_segment_commute: "closed_segment a b = closed_segment b a"
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Fri Sep 04 21:40:59 2015 +0200
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Fri Sep 04 22:16:54 2015 +0200
@@ -138,7 +138,7 @@
 qed
 
 lemma DERIV_caratheodory_within:
-  "(f has_field_derivative l) (at x within s) \<longleftrightarrow> 
+  "(f has_field_derivative l) (at x within s) \<longleftrightarrow>
    (\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> continuous (at x within s) g \<and> g x = l)"
       (is "?lhs = ?rhs")
 proof
@@ -209,6 +209,15 @@
   using has_derivative_at_within
   by blast
 
+lemma differentiable_at_imp_differentiable_on:
+  "(\<And>x. x \<in> s \<Longrightarrow> f differentiable at x) \<Longrightarrow> f differentiable_on s"
+  by (metis differentiable_at_withinI differentiable_on_def)
+
+corollary differentiable_iff_scaleR:
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  shows "f differentiable F \<longleftrightarrow> (\<exists>d. (f has_derivative (\<lambda>x. x *\<^sub>R d)) F)"
+  by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR)
+
 lemma differentiable_within_open: (* TODO: delete *)
   assumes "a \<in> s"
     and "open s"
@@ -2241,6 +2250,24 @@
   apply auto
   done
 
+lemma has_vector_derivative_id_at [simp]: "vector_derivative (\<lambda>x. x) (at a) = 1"
+  by (simp add: vector_derivative_at)
+
+lemma vector_derivative_minus_at [simp]:
+  "f differentiable at a
+   \<Longrightarrow> vector_derivative (\<lambda>x. - f x) (at a) = - vector_derivative f (at a)"
+  by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric])
+
+lemma vector_derivative_add_at [simp]:
+  "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk>
+   \<Longrightarrow> vector_derivative (\<lambda>x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)"
+  by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric])
+
+lemma vector_derivative_diff_at [simp]:
+  "\<lbrakk>f differentiable at a; g differentiable at a\<rbrakk>
+   \<Longrightarrow> vector_derivative (\<lambda>x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)"
+  by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric])
+
 lemma vector_derivative_within_closed_interval:
   assumes "a < b"
     and "x \<in> cbox a b"