HOL-Multivariate_Analysis: rename theories for more descriptive names
authorhoelzl
Thu, 04 Aug 2016 19:36:31 +0200
changeset 63594 bd218a9320b5
parent 63593 bbcb05504fdc
child 63595 aca2659ebba7
HOL-Multivariate_Analysis: rename theories for more descriptive names
src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy
src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem.thy
src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy
src/HOL/Multivariate_Analysis/Complex_Transcendental.thy
src/HOL/Multivariate_Analysis/Conformal_Mappings.thy
src/HOL/Multivariate_Analysis/Continuous_Extension.thy
src/HOL/Multivariate_Analysis/Extension.thy
src/HOL/Multivariate_Analysis/Fashoda.thy
src/HOL/Multivariate_Analysis/Fashoda_Theorem.thy
src/HOL/Multivariate_Analysis/Gamma.thy
src/HOL/Multivariate_Analysis/Gamma_Function.thy
src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy
src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy
src/HOL/Multivariate_Analysis/Henstock_Kurzweil_Integration.thy
src/HOL/Multivariate_Analysis/Integral_Test.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy
src/HOL/Multivariate_Analysis/Path_Connected.thy
src/HOL/Multivariate_Analysis/PolyRoots.thy
src/HOL/Multivariate_Analysis/Poly_Roots.thy
src/HOL/Multivariate_Analysis/Summation.thy
src/HOL/Multivariate_Analysis/Summation_Tests.thy
src/HOL/Multivariate_Analysis/Uniform_Limit.thy
src/HOL/Multivariate_Analysis/Weierstrass.thy
src/HOL/Multivariate_Analysis/Weierstrass_Theorems.thy
src/HOL/ROOT
--- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -1,7 +1,7 @@
 section \<open>Bounded Continuous Functions\<close>
 
 theory Bounded_Continuous_Function
-imports Integration
+imports Henstock_Kurzweil_Integration
 begin
 
 subsection \<open>Definition\<close>
--- a/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -1,7 +1,7 @@
 section \<open>Instantiates the finite Cartesian product of Euclidean spaces as a Euclidean space.\<close>
 
 theory Cartesian_Euclidean_Space
-imports Finite_Cartesian_Product Integration
+imports Finite_Cartesian_Product Henstock_Kurzweil_Integration
 begin
 
 lemma subspace_special_hyperplane: "subspace {x. x $ k = 0}"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,7539 @@
+section \<open>Complex path integrals and Cauchy's integral theorem\<close>
+
+text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
+
+theory Cauchy_Integral_Theorem
+imports Complex_Transcendental Weierstrass_Theorems Ordered_Euclidean_Space
+begin
+
+subsection\<open>Homeomorphisms of arc images\<close>
+
+lemma homeomorphism_arc:
+  fixes g :: "real \<Rightarrow> 'a::t2_space"
+  assumes "arc g"
+  obtains h where "homeomorphism {0..1} (path_image g) g h"
+using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def)
+
+lemma homeomorphic_arc_image_interval:
+  fixes g :: "real \<Rightarrow> 'a::t2_space" and a::real
+  assumes "arc g" "a < b"
+  shows "(path_image g) homeomorphic {a..b}"
+proof -
+  have "(path_image g) homeomorphic {0..1::real}"
+    by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc)
+  also have "... homeomorphic {a..b}"
+    using assms by (force intro: homeomorphic_closed_intervals_real)
+  finally show ?thesis .
+qed
+
+lemma homeomorphic_arc_images:
+  fixes g :: "real \<Rightarrow> 'a::t2_space" and h :: "real \<Rightarrow> 'b::t2_space"
+  assumes "arc g" "arc h"
+  shows "(path_image g) homeomorphic (path_image h)"
+proof -
+  have "(path_image g) homeomorphic {0..1::real}"
+    by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc)
+  also have "... homeomorphic (path_image h)"
+    by (meson assms homeomorphic_def homeomorphism_arc)
+  finally show ?thesis .
+qed
+
+subsection \<open>Piecewise differentiable functions\<close>
+
+definition piecewise_differentiable_on
+           (infixr "piecewise'_differentiable'_on" 50)
+  where "f piecewise_differentiable_on i  \<equiv>
+           continuous_on i f \<and>
+           (\<exists>s. finite s \<and> (\<forall>x \<in> i - s. f differentiable (at x within i)))"
+
+lemma piecewise_differentiable_on_imp_continuous_on:
+    "f piecewise_differentiable_on s \<Longrightarrow> continuous_on s f"
+by (simp add: piecewise_differentiable_on_def)
+
+lemma piecewise_differentiable_on_subset:
+    "f piecewise_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_differentiable_on t"
+  using continuous_on_subset
+  unfolding piecewise_differentiable_on_def
+  apply safe
+  apply (blast intro: elim: continuous_on_subset)
+  by (meson Diff_iff differentiable_within_subset subsetCE)
+
+lemma differentiable_on_imp_piecewise_differentiable:
+  fixes a:: "'a::{linorder_topology,real_normed_vector}"
+  shows "f differentiable_on {a..b} \<Longrightarrow> f piecewise_differentiable_on {a..b}"
+  apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on)
+  apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def)
+  done
+
+lemma differentiable_imp_piecewise_differentiable:
+    "(\<And>x. x \<in> s \<Longrightarrow> f differentiable (at x within s))
+         \<Longrightarrow> f piecewise_differentiable_on s"
+by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def
+         intro: differentiable_within_subset)
+
+lemma piecewise_differentiable_const [iff]: "(\<lambda>x. z) piecewise_differentiable_on s"
+  by (simp add: differentiable_imp_piecewise_differentiable)
+
+lemma piecewise_differentiable_compose:
+    "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s);
+      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
+      \<Longrightarrow> (g o f) piecewise_differentiable_on s"
+  apply (simp add: piecewise_differentiable_on_def, safe)
+  apply (blast intro: continuous_on_compose2)
+  apply (rename_tac A B)
+  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
+  apply (blast intro: differentiable_chain_within)
+  done
+
+lemma piecewise_differentiable_affine:
+  fixes m::real
+  assumes "f piecewise_differentiable_on ((\<lambda>x. m *\<^sub>R x + c) ` s)"
+  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_differentiable_on s"
+proof (cases "m = 0")
+  case True
+  then show ?thesis
+    unfolding o_def
+    by (force intro: differentiable_imp_piecewise_differentiable differentiable_const)
+next
+  case False
+  show ?thesis
+    apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable])
+    apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+
+    done
+qed
+
+lemma piecewise_differentiable_cases:
+  fixes c::real
+  assumes "f piecewise_differentiable_on {a..c}"
+          "g piecewise_differentiable_on {c..b}"
+           "a \<le> c" "c \<le> b" "f c = g c"
+  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_differentiable_on {a..b}"
+proof -
+  obtain s t where st: "finite s" "finite t"
+                       "\<forall>x\<in>{a..c} - s. f differentiable at x within {a..c}"
+                       "\<forall>x\<in>{c..b} - t. g differentiable at x within {c..b}"
+    using assms
+    by (auto simp: piecewise_differentiable_on_def)
+  have finabc: "finite ({a,b,c} \<union> (s \<union> t))"
+    by (metis \<open>finite s\<close> \<open>finite t\<close> finite_Un finite_insert finite.emptyI)
+  have "continuous_on {a..c} f" "continuous_on {c..b} g"
+    using assms piecewise_differentiable_on_def by auto
+  then have "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
+    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
+                               OF closed_real_atLeastAtMost [of c b],
+                               of f g "\<lambda>x. x\<le>c"]  assms
+    by (force simp: ivl_disj_un_two_touch)
+  moreover
+  { fix x
+    assume x: "x \<in> {a..b} - ({a,b,c} \<union> (s \<union> t))"
+    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg")
+    proof (cases x c rule: le_cases)
+      case le show ?diff_fg
+        apply (rule differentiable_transform_within [where d = "dist x c" and f = f])
+        using x le st
+        apply (simp_all add: dist_real_def)
+        apply (rule differentiable_at_withinI)
+        apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all)
+        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
+        apply (force elim!: differentiable_subset)+
+        done
+    next
+      case ge show ?diff_fg
+        apply (rule differentiable_transform_within [where d = "dist x c" and f = g])
+        using x ge st
+        apply (simp_all add: dist_real_def)
+        apply (rule differentiable_at_withinI)
+        apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all)
+        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
+        apply (force elim!: differentiable_subset)+
+        done
+    qed
+  }
+  then have "\<exists>s. finite s \<and>
+                 (\<forall>x\<in>{a..b} - s. (\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b})"
+    by (meson finabc)
+  ultimately show ?thesis
+    by (simp add: piecewise_differentiable_on_def)
+qed
+
+lemma piecewise_differentiable_neg:
+    "f piecewise_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_differentiable_on s"
+  by (auto simp: piecewise_differentiable_on_def continuous_on_minus)
+
+lemma piecewise_differentiable_add:
+  assumes "f piecewise_differentiable_on i"
+          "g piecewise_differentiable_on i"
+    shows "(\<lambda>x. f x + g x) piecewise_differentiable_on i"
+proof -
+  obtain s t where st: "finite s" "finite t"
+                       "\<forall>x\<in>i - s. f differentiable at x within i"
+                       "\<forall>x\<in>i - t. g differentiable at x within i"
+    using assms by (auto simp: piecewise_differentiable_on_def)
+  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)"
+    by auto
+  moreover have "continuous_on i f" "continuous_on i g"
+    using assms piecewise_differentiable_on_def by auto
+  ultimately show ?thesis
+    by (auto simp: piecewise_differentiable_on_def continuous_on_add)
+qed
+
+lemma piecewise_differentiable_diff:
+    "\<lbrakk>f piecewise_differentiable_on s;  g piecewise_differentiable_on s\<rbrakk>
+     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_differentiable_on s"
+  unfolding diff_conv_add_uminus
+  by (metis piecewise_differentiable_add piecewise_differentiable_neg)
+
+lemma continuous_on_joinpaths_D1:
+    "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1"
+  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"])
+  apply (rule continuous_intros | simp)+
+  apply (auto elim!: continuous_on_subset simp: joinpaths_def)
+  done
+
+lemma continuous_on_joinpaths_D2:
+    "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2"
+  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\<lambda>x. inverse 2*x + 1/2)"])
+  apply (rule continuous_intros | simp)+
+  apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def)
+  done
+
+lemma piecewise_differentiable_D1:
+    "(g1 +++ g2) piecewise_differentiable_on {0..1} \<Longrightarrow> g1 piecewise_differentiable_on {0..1}"
+  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1)
+  apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
+  apply simp
+  apply (intro ballI)
+  apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))"
+       in differentiable_transform_within)
+  apply (auto simp: dist_real_def joinpaths_def)
+  apply (rule differentiable_chain_within derivative_intros | simp)+
+  apply (rule differentiable_subset)
+  apply (force simp:)+
+  done
+
+lemma piecewise_differentiable_D2:
+    "\<lbrakk>(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\<rbrakk>
+    \<Longrightarrow> g2 piecewise_differentiable_on {0..1}"
+  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2)
+  apply (rule_tac x="insert 0 ((\<lambda>x. 2*x-1)`s)" in exI)
+  apply simp
+  apply (intro ballI)
+  apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)"
+          in differentiable_transform_within)
+  apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm)
+  apply (rule differentiable_chain_within derivative_intros | simp)+
+  apply (rule differentiable_subset)
+  apply (force simp: divide_simps)+
+  done
+
+
+subsubsection\<open>The concept of continuously differentiable\<close>
+
+text \<open>
+John Harrison writes as follows:
+
+``The usual assumption in complex analysis texts is that a path \<open>\<gamma>\<close> should be piecewise
+continuously differentiable, which ensures that the path integral exists at least for any continuous
+f, since all piecewise continuous functions are integrable. However, our notion of validity is
+weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a
+finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to
+the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this
+can integrate all derivatives.''
+
+"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec.
+Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165.
+
+And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably
+difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem
+asserting that all derivatives can be integrated before we can adopt Harrison's choice.\<close>
+
+definition C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
+           (infix "C1'_differentiable'_on" 50)
+  where
+  "f C1_differentiable_on s \<longleftrightarrow>
+   (\<exists>D. (\<forall>x \<in> s. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on s D)"
+
+lemma C1_differentiable_on_eq:
+    "f C1_differentiable_on s \<longleftrightarrow>
+     (\<forall>x \<in> s. f differentiable at x) \<and> continuous_on s (\<lambda>x. vector_derivative f (at x))"
+  unfolding C1_differentiable_on_def
+  apply safe
+  using differentiable_def has_vector_derivative_def apply blast
+  apply (erule continuous_on_eq)
+  using vector_derivative_at apply fastforce
+  using vector_derivative_works apply fastforce
+  done
+
+lemma C1_differentiable_on_subset:
+  "f C1_differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f C1_differentiable_on s"
+  unfolding C1_differentiable_on_def  continuous_on_eq_continuous_within
+  by (blast intro:  continuous_within_subset)
+
+lemma C1_differentiable_compose:
+    "\<lbrakk>f C1_differentiable_on s; g C1_differentiable_on (f ` s);
+      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
+      \<Longrightarrow> (g o f) C1_differentiable_on s"
+  apply (simp add: C1_differentiable_on_eq, safe)
+   using differentiable_chain_at apply blast
+  apply (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"])
+   apply (rule Limits.continuous_on_scaleR, assumption)
+   apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def)
+  by (simp add: vector_derivative_chain_at)
+
+lemma C1_diff_imp_diff: "f C1_differentiable_on s \<Longrightarrow> f differentiable_on s"
+  by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on)
+
+lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\<lambda>x. x) C1_differentiable_on s"
+  by (auto simp: C1_differentiable_on_eq continuous_on_const)
+
+lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on s"
+  by (auto simp: C1_differentiable_on_eq continuous_on_const)
+
+lemma C1_differentiable_on_add [simp, derivative_intros]:
+  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x + g x) C1_differentiable_on s"
+  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
+
+lemma C1_differentiable_on_minus [simp, derivative_intros]:
+  "f C1_differentiable_on s \<Longrightarrow> (\<lambda>x. - f x) C1_differentiable_on s"
+  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
+
+lemma C1_differentiable_on_diff [simp, derivative_intros]:
+  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x - g x) C1_differentiable_on s"
+  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
+
+lemma C1_differentiable_on_mult [simp, derivative_intros]:
+  fixes f g :: "real \<Rightarrow> 'a :: real_normed_algebra"
+  shows "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x * g x) C1_differentiable_on s"
+  unfolding C1_differentiable_on_eq
+  by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within)
+
+lemma C1_differentiable_on_scaleR [simp, derivative_intros]:
+  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) C1_differentiable_on s"
+  unfolding C1_differentiable_on_eq
+  by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+
+
+
+definition piecewise_C1_differentiable_on
+           (infixr "piecewise'_C1'_differentiable'_on" 50)
+  where "f piecewise_C1_differentiable_on i  \<equiv>
+           continuous_on i f \<and>
+           (\<exists>s. finite s \<and> (f C1_differentiable_on (i - s)))"
+
+lemma C1_differentiable_imp_piecewise:
+    "f C1_differentiable_on s \<Longrightarrow> f piecewise_C1_differentiable_on s"
+  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within)
+
+lemma piecewise_C1_imp_differentiable:
+    "f piecewise_C1_differentiable_on i \<Longrightarrow> f piecewise_differentiable_on i"
+  by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def
+           C1_differentiable_on_def differentiable_def has_vector_derivative_def
+           intro: has_derivative_at_within)
+
+lemma piecewise_C1_differentiable_compose:
+    "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s);
+      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
+      \<Longrightarrow> (g o f) piecewise_C1_differentiable_on s"
+  apply (simp add: piecewise_C1_differentiable_on_def, safe)
+  apply (blast intro: continuous_on_compose2)
+  apply (rename_tac A B)
+  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
+  apply (rule conjI, blast)
+  apply (rule C1_differentiable_compose)
+  apply (blast intro: C1_differentiable_on_subset)
+  apply (blast intro: C1_differentiable_on_subset)
+  by (simp add: Diff_Int_distrib2)
+
+lemma piecewise_C1_differentiable_on_subset:
+    "f piecewise_C1_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_C1_differentiable_on t"
+  by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset)
+
+lemma C1_differentiable_imp_continuous_on:
+  "f C1_differentiable_on s \<Longrightarrow> continuous_on s f"
+  unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within
+  using differentiable_at_withinI differentiable_imp_continuous_within by blast
+
+lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}"
+  unfolding C1_differentiable_on_def
+  by auto
+
+lemma piecewise_C1_differentiable_affine:
+  fixes m::real
+  assumes "f piecewise_C1_differentiable_on ((\<lambda>x. m * x + c) ` s)"
+  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s"
+proof (cases "m = 0")
+  case True
+  then show ?thesis
+    unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const)
+next
+  case False
+  show ?thesis
+    apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise])
+    apply (rule assms derivative_intros | simp add: False vimage_def)+
+    using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real]
+    apply simp
+    done
+qed
+
+lemma piecewise_C1_differentiable_cases:
+  fixes c::real
+  assumes "f piecewise_C1_differentiable_on {a..c}"
+          "g piecewise_C1_differentiable_on {c..b}"
+           "a \<le> c" "c \<le> b" "f c = g c"
+  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_C1_differentiable_on {a..b}"
+proof -
+  obtain s t where st: "f C1_differentiable_on ({a..c} - s)"
+                       "g C1_differentiable_on ({c..b} - t)"
+                       "finite s" "finite t"
+    using assms
+    by (force simp: piecewise_C1_differentiable_on_def)
+  then have f_diff: "f differentiable_on {a..<c} - s"
+        and g_diff: "g differentiable_on {c<..b} - t"
+    by (simp_all add: C1_differentiable_on_eq differentiable_at_withinI differentiable_on_def)
+  have "continuous_on {a..c} f" "continuous_on {c..b} g"
+    using assms piecewise_C1_differentiable_on_def by auto
+  then have cab: "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
+    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
+                               OF closed_real_atLeastAtMost [of c b],
+                               of f g "\<lambda>x. x\<le>c"]  assms
+    by (force simp: ivl_disj_un_two_touch)
+  { fix x
+    assume x: "x \<in> {a..b} - insert c (s \<union> t)"
+    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x" (is "?diff_fg")
+    proof (cases x c rule: le_cases)
+      case le show ?diff_fg
+        apply (rule differentiable_transform_within [where f=f and d = "dist x c"])
+        using x dist_real_def le st by (auto simp: C1_differentiable_on_eq)
+    next
+      case ge show ?diff_fg
+        apply (rule differentiable_transform_within [where f=g and d = "dist x c"])
+        using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq)
+    qed
+  }
+  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)"
+    by auto
+  moreover
+  { assume fcon: "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative f (at x))"
+       and gcon: "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative g (at x))"
+    have "open ({a<..<c} - s)"  "open ({c<..<b} - t)"
+      using st by (simp_all add: open_Diff finite_imp_closed)
+    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))"
+      apply (rule continuous_on_eq [OF fcon])
+      apply (simp add:)
+      apply (rule vector_derivative_at [symmetric])
+      apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within)
+      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
+      apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1))
+      apply auto
+      done
+    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))"
+      apply (rule continuous_on_eq [OF gcon])
+      apply (simp add:)
+      apply (rule vector_derivative_at [symmetric])
+      apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within)
+      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
+      apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2))
+      apply auto
+      done
+    ultimately 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))"
+      apply (rule continuous_on_subset [OF continuous_on_open_Un], auto)
+      done
+  } note * = this
+  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))"
+    using st
+    by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *)
+  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)"
+    apply (rule_tac x="{a,b,c} \<union> s \<union> t" in exI)
+    using st  by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset)
+  with cab show ?thesis
+    by (simp add: piecewise_C1_differentiable_on_def)
+qed
+
+lemma piecewise_C1_differentiable_neg:
+    "f piecewise_C1_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_C1_differentiable_on s"
+  unfolding piecewise_C1_differentiable_on_def
+  by (auto intro!: continuous_on_minus C1_differentiable_on_minus)
+
+lemma piecewise_C1_differentiable_add:
+  assumes "f piecewise_C1_differentiable_on i"
+          "g piecewise_C1_differentiable_on i"
+    shows "(\<lambda>x. f x + g x) piecewise_C1_differentiable_on i"
+proof -
+  obtain s t where st: "finite s" "finite t"
+                       "f C1_differentiable_on (i-s)"
+                       "g C1_differentiable_on (i-t)"
+    using assms by (auto simp: piecewise_C1_differentiable_on_def)
+  then have "finite (s \<union> t) \<and> (\<lambda>x. f x + g x) C1_differentiable_on i - (s \<union> t)"
+    by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset)
+  moreover have "continuous_on i f" "continuous_on i g"
+    using assms piecewise_C1_differentiable_on_def by auto
+  ultimately show ?thesis
+    by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add)
+qed
+
+lemma piecewise_C1_differentiable_diff:
+    "\<lbrakk>f piecewise_C1_differentiable_on s;  g piecewise_C1_differentiable_on s\<rbrakk>
+     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_C1_differentiable_on s"
+  unfolding diff_conv_add_uminus
+  by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg)
+
+lemma piecewise_C1_differentiable_D1:
+  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
+  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}"
+    shows "g1 piecewise_C1_differentiable_on {0..1}"
+proof -
+  obtain s where "finite s"
+             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
+             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
+    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  then have g1D: "g1 differentiable at x" if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
+    apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within)
+    using that
+    apply (simp_all add: dist_real_def joinpaths_def)
+    apply (rule differentiable_chain_at derivative_intros | force)+
+    done
+  have [simp]: "vector_derivative (g1 \<circ> op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)"
+               if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
+    apply (subst vector_derivative_chain_at)
+    using that
+    apply (rule derivative_eq_intros g1D | simp)+
+    done
+  have "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
+    using co12 by (rule continuous_on_subset) force
+  then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 o op*2) (at x))"
+    apply (rule continuous_on_eq [OF _ vector_derivative_at])
+    apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within)
+    apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric])
+    apply (force intro: g1D differentiable_chain_at)
+    apply auto
+    done
+  have "continuous_on ({0..1} - insert 1 (op * 2 ` s))
+                      ((\<lambda>x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))"
+    apply (rule continuous_intros)+
+    using coDhalf
+    apply (simp add: scaleR_conv_of_real image_set_diff image_image)
+    done
+  then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\<lambda>x. vector_derivative g1 (at x))"
+    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
+  have "continuous_on {0..1} g1"
+    using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast
+  with \<open>finite s\<close> show ?thesis
+    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+    apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
+    apply (simp add: g1D con_g1)
+  done
+qed
+
+lemma piecewise_C1_differentiable_D2:
+  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
+  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2"
+    shows "g2 piecewise_C1_differentiable_on {0..1}"
+proof -
+  obtain s where "finite s"
+             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
+             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
+    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  then have g2D: "g2 differentiable at x" if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
+    apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" in differentiable_transform_within)
+    using that
+    apply (simp_all add: dist_real_def joinpaths_def)
+    apply (auto simp: dist_real_def joinpaths_def field_simps)
+    apply (rule differentiable_chain_at derivative_intros | force)+
+    apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps)
+    apply assumption
+    done
+  have [simp]: "vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)"
+               if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
+    using that  by (auto simp: vector_derivative_chain_at divide_simps g2D)
+  have "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
+    using co12 by (rule continuous_on_subset) force
+  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))"
+    apply (rule continuous_on_eq [OF _ vector_derivative_at])
+    apply (rule_tac f="g2 o (\<lambda>x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within)
+    apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric]
+                intro!: g2D differentiable_chain_at)
+    done
+  have [simp]: "((\<lambda>x. (x + 1) / 2) ` ({0..1} - insert 0 ((\<lambda>x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)"
+    apply (simp add: image_set_diff inj_on_def image_image)
+    apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib)
+    done
+  have "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s))
+                      ((\<lambda>x. 1/2 * vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at x)) o (\<lambda>x. (x+1)/2))"
+    by (rule continuous_intros | simp add:  coDhalf)+
+  then have con_g2: "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) (\<lambda>x. vector_derivative g2 (at x))"
+    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
+  have "continuous_on {0..1} g2"
+    using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast
+  with \<open>finite s\<close> show ?thesis
+    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+    apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` s)" in exI)
+    apply (simp add: g2D con_g2)
+  done
+qed
+
+subsection \<open>Valid paths, and their start and finish\<close>
+
+lemma Diff_Un_eq: "A - (B \<union> C) = A - B - C"
+  by blast
+
+definition valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
+  where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}"
+
+definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
+  where "closed_path g \<equiv> g 0 = g 1"
+
+subsubsection\<open>In particular, all results for paths apply\<close>
+
+lemma valid_path_imp_path: "valid_path g \<Longrightarrow> path g"
+by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def)
+
+lemma connected_valid_path_image: "valid_path g \<Longrightarrow> connected(path_image g)"
+  by (metis connected_path_image valid_path_imp_path)
+
+lemma compact_valid_path_image: "valid_path g \<Longrightarrow> compact(path_image g)"
+  by (metis compact_path_image valid_path_imp_path)
+
+lemma bounded_valid_path_image: "valid_path g \<Longrightarrow> bounded(path_image g)"
+  by (metis bounded_path_image valid_path_imp_path)
+
+lemma closed_valid_path_image: "valid_path g \<Longrightarrow> closed(path_image g)"
+  by (metis closed_path_image valid_path_imp_path)
+
+proposition valid_path_compose:
+  assumes "valid_path g"
+      and der: "\<And>x. x \<in> path_image g \<Longrightarrow> \<exists>f'. (f has_field_derivative f') (at x)"
+      and con: "continuous_on (path_image g) (deriv f)"
+    shows "valid_path (f o g)"
+proof -
+  obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s"
+    using \<open>valid_path g\<close> unfolding valid_path_def piecewise_C1_differentiable_on_def by auto
+  have "f \<circ> g differentiable at t" when "t\<in>{0..1} - s" for t
+    proof (rule differentiable_chain_at)
+      show "g differentiable at t" using \<open>valid_path g\<close>
+        by (meson C1_differentiable_on_eq \<open>g C1_differentiable_on {0..1} - s\<close> that)
+    next
+      have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
+      then obtain f' where "(f has_field_derivative f') (at (g t))"
+        using der by auto
+      then have " (f has_derivative op * f') (at (g t))"
+        using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto
+      then show "f differentiable at (g t)" using differentiableI by auto
+    qed
+  moreover have "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (f \<circ> g) (at x))"
+    proof (rule continuous_on_eq [where f = "\<lambda>x. vector_derivative g (at x) * deriv f (g x)"],
+        rule continuous_intros)
+      show "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative g (at x))"
+        using g_diff C1_differentiable_on_eq by auto
+    next
+      have "continuous_on {0..1} (\<lambda>x. deriv f (g x))"
+        using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def]
+          \<open>valid_path g\<close> piecewise_C1_differentiable_on_def valid_path_def
+        by blast
+      then show "continuous_on ({0..1} - s) (\<lambda>x. deriv f (g x))"
+        using continuous_on_subset by blast
+    next
+      show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \<circ> g) (at t)"
+          when "t \<in> {0..1} - s" for t
+        proof (rule vector_derivative_chain_at_general[symmetric])
+          show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that)
+        next
+          have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
+          then obtain f' where "(f has_field_derivative f') (at (g t))"
+            using der by auto
+          then show "\<exists>g'. (f has_field_derivative g') (at (g t))" by auto
+        qed
+    qed
+  ultimately have "f o g C1_differentiable_on {0..1} - s"
+    using C1_differentiable_on_eq by blast
+  moreover have "path (f o g)"
+    proof -
+      have "isCont f x" when "x\<in>path_image g" for x
+        proof -
+          obtain f' where "(f has_field_derivative f') (at x)"
+            using der[rule_format] \<open>x\<in>path_image g\<close> by auto
+          thus ?thesis using DERIV_isCont by auto
+        qed
+      then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto
+      then show ?thesis using path_continuous_image \<open>valid_path g\<close> valid_path_imp_path by auto
+    qed
+  ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def
+    using \<open>finite s\<close> by auto
+qed
+
+
+subsection\<open>Contour Integrals along a path\<close>
+
+text\<open>This definition is for complex numbers only, and does not generalise to line integrals in a vector field\<close>
+
+text\<open>piecewise differentiable function on [0,1]\<close>
+
+definition has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
+           (infixr "has'_contour'_integral" 50)
+  where "(f has_contour_integral i) g \<equiv>
+           ((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
+            has_integral i) {0..1}"
+
+definition contour_integrable_on
+           (infixr "contour'_integrable'_on" 50)
+  where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
+
+definition contour_integral
+  where "contour_integral g f \<equiv> @i. (f has_contour_integral i) g \<or> ~ f contour_integrable_on g \<and> i=0"
+
+lemma not_integrable_contour_integral: "~ f contour_integrable_on g \<Longrightarrow> contour_integral g f = 0"
+  unfolding contour_integrable_on_def contour_integral_def by blast
+
+lemma contour_integral_unique: "(f has_contour_integral i) g \<Longrightarrow> contour_integral g f = i"
+  apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def)
+  using has_integral_unique by blast
+
+corollary has_contour_integral_eqpath:
+     "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>;
+       contour_integral p f = contour_integral \<gamma> f\<rbrakk>
+      \<Longrightarrow> (f has_contour_integral y) \<gamma>"
+using contour_integrable_on_def contour_integral_unique by auto
+
+lemma has_contour_integral_integral:
+    "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i"
+  by (metis contour_integral_unique contour_integrable_on_def)
+
+lemma has_contour_integral_unique:
+    "(f has_contour_integral i) g \<Longrightarrow> (f has_contour_integral j) g \<Longrightarrow> i = j"
+  using has_integral_unique
+  by (auto simp: has_contour_integral_def)
+
+lemma has_contour_integral_integrable: "(f has_contour_integral i) g \<Longrightarrow> f contour_integrable_on g"
+  using contour_integrable_on_def by blast
+
+(* Show that we can forget about the localized derivative.*)
+
+lemma vector_derivative_within_interior:
+     "\<lbrakk>x \<in> interior s; NO_MATCH UNIV s\<rbrakk>
+      \<Longrightarrow> vector_derivative f (at x within s) = vector_derivative f (at x)"
+  apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior)
+  apply (subst lim_within_interior, auto)
+  done
+
+lemma has_integral_localized_vector_derivative:
+    "((\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \<longleftrightarrow>
+     ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}"
+proof -
+  have "{a..b} - {a,b} = interior {a..b}"
+    by (simp add: atLeastAtMost_diff_ends)
+  show ?thesis
+    apply (rule has_integral_spike_eq [of "{a,b}"])
+    apply (auto simp: vector_derivative_within_interior)
+    done
+qed
+
+lemma integrable_on_localized_vector_derivative:
+    "(\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \<longleftrightarrow>
+     (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}"
+  by (simp add: integrable_on_def has_integral_localized_vector_derivative)
+
+lemma has_contour_integral:
+     "(f has_contour_integral i) g \<longleftrightarrow>
+      ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
+  by (simp add: has_integral_localized_vector_derivative has_contour_integral_def)
+
+lemma contour_integrable_on:
+     "f contour_integrable_on g \<longleftrightarrow>
+      (\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}"
+  by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def)
+
+subsection\<open>Reversing a path\<close>
+
+lemma valid_path_imp_reverse:
+  assumes "valid_path g"
+    shows "valid_path(reversepath g)"
+proof -
+  obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
+  then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))"
+    apply (auto simp: reversepath_def)
+    apply (rule C1_differentiable_compose [of "\<lambda>x::real. 1-x" _ g, unfolded o_def])
+    apply (auto simp: C1_differentiable_on_eq)
+    apply (rule continuous_intros, force)
+    apply (force elim!: continuous_on_subset)
+    apply (simp add: finite_vimageI inj_on_def)
+    done
+  then show ?thesis using assms
+    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric])
+qed
+
+lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \<longleftrightarrow> valid_path g"
+  using valid_path_imp_reverse by force
+
+lemma has_contour_integral_reversepath:
+  assumes "valid_path g" "(f has_contour_integral i) g"
+    shows "(f has_contour_integral (-i)) (reversepath g)"
+proof -
+  { fix s x
+    assume xs: "g C1_differentiable_on ({0..1} - s)" "x \<notin> op - 1 ` s" "0 \<le> x" "x \<le> 1"
+      have "vector_derivative (\<lambda>x. g (1 - x)) (at x within {0..1}) =
+            - vector_derivative g (at (1 - x) within {0..1})"
+      proof -
+        obtain f' where f': "(g has_vector_derivative f') (at (1 - x))"
+          using xs
+          by (force simp: has_vector_derivative_def C1_differentiable_on_def)
+        have "(g o (\<lambda>x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)"
+          apply (rule vector_diff_chain_within)
+          apply (intro vector_diff_chain_within derivative_eq_intros | simp)+
+          apply (rule has_vector_derivative_at_within [OF f'])
+          done
+        then have mf': "((\<lambda>x. g (1 - x)) has_vector_derivative -f') (at x)"
+          by (simp add: o_def)
+        show ?thesis
+          using xs
+          by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f'])
+      qed
+  } note * = this
+  have 01: "{0..1::real} = cbox 0 1"
+    by simp
+  show ?thesis using assms
+    apply (auto simp: has_contour_integral_def)
+    apply (drule has_integral_affinity01 [where m= "-1" and c=1])
+    apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def)
+    apply (drule has_integral_neg)
+    apply (rule_tac s = "(\<lambda>x. 1 - x) ` s" in has_integral_spike_finite)
+    apply (auto simp: *)
+    done
+qed
+
+lemma contour_integrable_reversepath:
+    "valid_path g \<Longrightarrow> f contour_integrable_on g \<Longrightarrow> f contour_integrable_on (reversepath g)"
+  using has_contour_integral_reversepath contour_integrable_on_def by blast
+
+lemma contour_integrable_reversepath_eq:
+    "valid_path g \<Longrightarrow> (f contour_integrable_on (reversepath g) \<longleftrightarrow> f contour_integrable_on g)"
+  using contour_integrable_reversepath valid_path_reversepath by fastforce
+
+lemma contour_integral_reversepath:
+  assumes "valid_path g"
+    shows "contour_integral (reversepath g) f = - (contour_integral g f)"
+proof (cases "f contour_integrable_on g")
+  case True then show ?thesis
+    by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath)
+next
+  case False then have "~ f contour_integrable_on (reversepath g)"
+    by (simp add: assms contour_integrable_reversepath_eq)
+  with False show ?thesis by (simp add: not_integrable_contour_integral)
+qed
+
+
+subsection\<open>Joining two paths together\<close>
+
+lemma valid_path_join:
+  assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2"
+    shows "valid_path(g1 +++ g2)"
+proof -
+  have "g1 1 = g2 0"
+    using assms by (auto simp: pathfinish_def pathstart_def)
+  moreover have "(g1 o (\<lambda>x. 2*x)) piecewise_C1_differentiable_on {0..1/2}"
+    apply (rule piecewise_C1_differentiable_compose)
+    using assms
+    apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths)
+    apply (rule continuous_intros | simp)+
+    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
+    done
+  moreover have "(g2 o (\<lambda>x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}"
+    apply (rule piecewise_C1_differentiable_compose)
+    using assms unfolding valid_path_def piecewise_C1_differentiable_on_def
+    by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI
+             simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths)
+  ultimately show ?thesis
+    apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def)
+    apply (rule piecewise_C1_differentiable_cases)
+    apply (auto simp: o_def)
+    done
+qed
+
+lemma valid_path_join_D1:
+  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
+  shows "valid_path (g1 +++ g2) \<Longrightarrow> valid_path g1"
+  unfolding valid_path_def
+  by (rule piecewise_C1_differentiable_D1)
+
+lemma valid_path_join_D2:
+  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
+  shows "\<lbrakk>valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> valid_path g2"
+  unfolding valid_path_def
+  by (rule piecewise_C1_differentiable_D2)
+
+lemma valid_path_join_eq [simp]:
+  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
+  shows "pathfinish g1 = pathstart g2 \<Longrightarrow> (valid_path(g1 +++ g2) \<longleftrightarrow> valid_path g1 \<and> valid_path g2)"
+  using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast
+
+lemma has_contour_integral_join:
+  assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2"
+          "valid_path g1" "valid_path g2"
+    shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)"
+proof -
+  obtain s1 s2
+    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
+      and s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
+    using assms
+    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  have 1: "((\<lambda>x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}"
+   and 2: "((\<lambda>x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}"
+    using assms
+    by (auto simp: has_contour_integral)
+  have i1: "((\<lambda>x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}"
+   and i2: "((\<lambda>x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}"
+    using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]]
+          has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]]
+    by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac)
+  have g1: "\<lbrakk>0 \<le> z; z*2 < 1; z*2 \<notin> s1\<rbrakk> \<Longrightarrow>
+            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
+            2 *\<^sub>R vector_derivative g1 (at (z*2))" for z
+    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>z - 1/2\<bar>"]])
+    apply (simp_all add: dist_real_def abs_if split: if_split_asm)
+    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x" 2 _ g1, simplified o_def])
+    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
+    using s1
+    apply (auto simp: algebra_simps vector_derivative_works)
+    done
+  have g2: "\<lbrakk>1 < z*2; z \<le> 1; z*2 - 1 \<notin> s2\<rbrakk> \<Longrightarrow>
+            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
+            2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z
+    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2 (2*x - 1))" and d = "\<bar>z - 1/2\<bar>"]])
+    apply (simp_all add: dist_real_def abs_if split: if_split_asm)
+    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x - 1" 2 _ g2, simplified o_def])
+    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
+    using s2
+    apply (auto simp: algebra_simps vector_derivative_works)
+    done
+  have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}"
+    apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"])
+    using s1
+    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
+    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1)
+    done
+  moreover have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}"
+    apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\<lambda>x. 2*x-1) -` s2)"])
+    using s2
+    apply (force intro: finite_vimageI [where h = "\<lambda>x. 2*x-1"] inj_onI)
+    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2)
+    done
+  ultimately
+  show ?thesis
+    apply (simp add: has_contour_integral)
+    apply (rule has_integral_combine [where c = "1/2"], auto)
+    done
+qed
+
+lemma contour_integrable_joinI:
+  assumes "f contour_integrable_on g1" "f contour_integrable_on g2"
+          "valid_path g1" "valid_path g2"
+    shows "f contour_integrable_on (g1 +++ g2)"
+  using assms
+  by (meson has_contour_integral_join contour_integrable_on_def)
+
+lemma contour_integrable_joinD1:
+  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1"
+    shows "f contour_integrable_on g1"
+proof -
+  obtain s1
+    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  have "(\<lambda>x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
+    using assms
+    apply (auto simp: contour_integrable_on)
+    apply (drule integrable_on_subcbox [where a=0 and b="1/2"])
+    apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified])
+    done
+  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
+    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
+  have g1: "\<lbrakk>0 < z; z < 1; z \<notin> s1\<rbrakk> \<Longrightarrow>
+            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) =
+            2 *\<^sub>R vector_derivative g1 (at z)"  for z
+    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>(z-1)/2\<bar>"]])
+    apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
+    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2" 2 _ g1, simplified o_def])
+    using s1
+    apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
+    done
+  show ?thesis
+    using s1
+    apply (auto simp: contour_integrable_on)
+    apply (rule integrable_spike_finite [of "{0,1} \<union> s1", OF _ _ *])
+    apply (auto simp: joinpaths_def scaleR_conv_of_real g1)
+    done
+qed
+
+lemma contour_integrable_joinD2:
+  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2"
+    shows "f contour_integrable_on g2"
+proof -
+  obtain s2
+    where s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  have "(\<lambda>x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}"
+    using assms
+    apply (auto simp: contour_integrable_on)
+    apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto)
+    apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified])
+    apply (simp add: image_affinity_atLeastAtMost_diff)
+    done
+  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2)))
+                integrable_on {0..1}"
+    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
+  have g2: "\<lbrakk>0 < z; z < 1; z \<notin> s2\<rbrakk> \<Longrightarrow>
+            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) =
+            2 *\<^sub>R vector_derivative g2 (at z)" for z
+    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2(2*x-1))" and d = "\<bar>z/2\<bar>"]])
+    apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
+    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2-1" 2 _ g2, simplified o_def])
+    using s2
+    apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left
+                      vector_derivative_works add_divide_distrib)
+    done
+  show ?thesis
+    using s2
+    apply (auto simp: contour_integrable_on)
+    apply (rule integrable_spike_finite [of "{0,1} \<union> s2", OF _ _ *])
+    apply (auto simp: joinpaths_def scaleR_conv_of_real g2)
+    done
+qed
+
+lemma contour_integrable_join [simp]:
+  shows
+    "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
+     \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2"
+using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
+
+lemma contour_integral_join [simp]:
+  shows
+    "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
+        \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f"
+  by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
+
+
+subsection\<open>Shifting the starting point of a (closed) path\<close>
+
+lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))"
+  by (auto simp: shiftpath_def)
+
+lemma valid_path_shiftpath [intro]:
+  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
+    shows "valid_path(shiftpath a g)"
+  using assms
+  apply (auto simp: valid_path_def shiftpath_alt_def)
+  apply (rule piecewise_C1_differentiable_cases)
+  apply (auto simp: algebra_simps)
+  apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
+  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
+  apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
+  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
+  done
+
+lemma has_contour_integral_shiftpath:
+  assumes f: "(f has_contour_integral i) g" "valid_path g"
+      and a: "a \<in> {0..1}"
+    shows "(f has_contour_integral i) (shiftpath a g)"
+proof -
+  obtain s
+    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  have *: "((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
+    using assms by (auto simp: has_contour_integral)
+  then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) +
+                    integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))"
+    apply (rule has_integral_unique)
+    apply (subst add.commute)
+    apply (subst integral_combine)
+    using assms * integral_unique by auto
+  { fix x
+    have "0 \<le> x \<Longrightarrow> x + a < 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a) ` s \<Longrightarrow>
+         vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))"
+      unfolding shiftpath_def
+      apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x))" and d = "dist(1-a) x"]])
+        apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
+      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a" 1 _ g, simplified o_def scaleR_one])
+       apply (intro derivative_eq_intros | simp)+
+      using g
+       apply (drule_tac x="x+a" in bspec)
+      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
+      done
+  } note vd1 = this
+  { fix x
+    have "1 < x + a \<Longrightarrow> x \<le> 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a + 1) ` s \<Longrightarrow>
+          vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))"
+      unfolding shiftpath_def
+      apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x-1))" and d = "dist (1-a) x"]])
+        apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
+      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a-1" 1 _ g, simplified o_def scaleR_one])
+       apply (intro derivative_eq_intros | simp)+
+      using g
+      apply (drule_tac x="x+a-1" in bspec)
+      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
+      done
+  } note vd2 = this
+  have va1: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})"
+    using * a   by (fastforce intro: integrable_subinterval_real)
+  have v0a: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})"
+    apply (rule integrable_subinterval_real)
+    using * a by auto
+  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
+        has_integral  integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {0..1 - a}"
+    apply (rule has_integral_spike_finite
+             [where s = "{1-a} \<union> (\<lambda>x. x-a) ` s" and f = "\<lambda>x. f(g(a+x)) * vector_derivative g (at(a+x))"])
+      using s apply blast
+     using a apply (auto simp: algebra_simps vd1)
+     apply (force simp: shiftpath_def add.commute)
+    using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]]
+    apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute)
+    done
+  moreover
+  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
+        has_integral  integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {1 - a..1}"
+    apply (rule has_integral_spike_finite
+             [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))"])
+      using s apply blast
+     using a apply (auto simp: algebra_simps vd2)
+     apply (force simp: shiftpath_def add.commute)
+    using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]]
+    apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified])
+    apply (simp add: algebra_simps)
+    done
+  ultimately show ?thesis
+    using a
+    by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"])
+qed
+
+lemma has_contour_integral_shiftpath_D:
+  assumes "(f has_contour_integral i) (shiftpath a g)"
+          "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
+    shows "(f has_contour_integral i) g"
+proof -
+  obtain s
+    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
+  { fix x
+    assume x: "0 < x" "x < 1" "x \<notin> s"
+    then have gx: "g differentiable at x"
+      using g by auto
+    have "vector_derivative g (at x within {0..1}) =
+          vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})"
+      apply (rule vector_derivative_at_within_ivl
+                  [OF has_vector_derivative_transform_within_open
+                      [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]])
+      using s g assms x
+      apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath
+                        vector_derivative_within_interior vector_derivative_works [symmetric])
+      apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"])
+      apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm)
+      done
+  } note vd = this
+  have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))"
+    using assms  by (auto intro!: has_contour_integral_shiftpath)
+  show ?thesis
+    apply (simp add: has_contour_integral_def)
+    apply (rule has_integral_spike_finite [of "{0,1} \<union> s", OF _ _  fi [unfolded has_contour_integral_def]])
+    using s assms vd
+    apply (auto simp: Path_Connected.shiftpath_shiftpath)
+    done
+qed
+
+lemma has_contour_integral_shiftpath_eq:
+  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
+    shows "(f has_contour_integral i) (shiftpath a g) \<longleftrightarrow> (f has_contour_integral i) g"
+  using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast
+
+lemma contour_integrable_on_shiftpath_eq:
+  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
+    shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g"
+using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto
+
+lemma contour_integral_shiftpath:
+  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
+    shows "contour_integral (shiftpath a g) f = contour_integral g f"
+   using assms
+   by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq)
+
+
+subsection\<open>More about straight-line paths\<close>
+
+lemma has_vector_derivative_linepath_within:
+    "(linepath a b has_vector_derivative (b - a)) (at x within s)"
+apply (simp add: linepath_def has_vector_derivative_def algebra_simps)
+apply (rule derivative_eq_intros | simp)+
+done
+
+lemma vector_derivative_linepath_within:
+    "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a"
+  apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified])
+  apply (auto simp: has_vector_derivative_linepath_within)
+  done
+
+lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a"
+  by (simp add: has_vector_derivative_linepath_within vector_derivative_at)
+
+lemma valid_path_linepath [iff]: "valid_path (linepath a b)"
+  apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath)
+  apply (rule_tac x="{}" in exI)
+  apply (simp add: differentiable_on_def differentiable_def)
+  using has_vector_derivative_def has_vector_derivative_linepath_within
+  apply (fastforce simp add: continuous_on_eq_continuous_within)
+  done
+
+lemma has_contour_integral_linepath:
+  shows "(f has_contour_integral i) (linepath a b) \<longleftrightarrow>
+         ((\<lambda>x. f(linepath a b x) * (b - a)) has_integral i) {0..1}"
+  by (simp add: has_contour_integral vector_derivative_linepath_at)
+
+lemma linepath_in_path:
+  shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b"
+  by (auto simp: segment linepath_def)
+
+lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b"
+  by (auto simp: segment linepath_def)
+
+lemma linepath_in_convex_hull:
+    fixes x::real
+    assumes a: "a \<in> convex hull s"
+        and b: "b \<in> convex hull s"
+        and x: "0\<le>x" "x\<le>1"
+       shows "linepath a b x \<in> convex hull s"
+  apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD])
+  using x
+  apply (auto simp: linepath_image_01 [symmetric])
+  done
+
+lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b"
+  by (simp add: linepath_def)
+
+lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0"
+  by (simp add: linepath_def)
+
+lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)"
+  by (simp add: has_contour_integral_linepath)
+
+lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0"
+  using has_contour_integral_trivial contour_integral_unique by blast
+
+
+subsection\<open>Relation to subpath construction\<close>
+
+lemma valid_path_subpath:
+  fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
+  assumes "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
+    shows "valid_path(subpath u v g)"
+proof (cases "v=u")
+  case True
+  then show ?thesis
+    unfolding valid_path_def subpath_def
+    by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise)
+next
+  case False
+  have "(g o (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}"
+    apply (rule piecewise_C1_differentiable_compose)
+    apply (simp add: C1_differentiable_imp_piecewise)
+     apply (simp add: image_affinity_atLeastAtMost)
+    using assms False
+    apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset)
+    apply (subst Int_commute)
+    apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI)
+    done
+  then show ?thesis
+    by (auto simp: o_def valid_path_def subpath_def)
+qed
+
+lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)"
+  by (simp add: has_contour_integral subpath_def)
+
+lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)"
+  using has_contour_integral_subpath_refl contour_integrable_on_def by blast
+
+lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0"
+  by (simp add: has_contour_integral_subpath_refl contour_integral_unique)
+
+lemma has_contour_integral_subpath:
+  assumes f: "f contour_integrable_on g" and g: "valid_path g"
+      and uv: "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
+    shows "(f has_contour_integral  integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x)))
+           (subpath u v g)"
+proof (cases "v=u")
+  case True
+  then show ?thesis
+    using f   by (simp add: contour_integrable_on_def subpath_def has_contour_integral)
+next
+  case False
+  obtain s where s: "\<And>x. x \<in> {0..1} - s \<Longrightarrow> g differentiable at x" and fs: "finite s"
+    using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast
+  have *: "((\<lambda>x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u)))
+            has_integral (1 / (v - u)) * integral {u..v} (\<lambda>t. f (g t) * vector_derivative g (at t)))
+           {0..1}"
+    using f uv
+    apply (simp add: contour_integrable_on subpath_def has_contour_integral)
+    apply (drule integrable_on_subcbox [where a=u and b=v, simplified])
+    apply (simp_all add: has_integral_integral)
+    apply (drule has_integral_affinity [where m="v-u" and c=u, simplified])
+    apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real)
+    apply (simp add: divide_simps False)
+    done
+  { fix x
+    have "x \<in> {0..1} \<Longrightarrow>
+           x \<notin> (\<lambda>t. (v-u) *\<^sub>R t + u) -` s \<Longrightarrow>
+           vector_derivative (\<lambda>x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))"
+      apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]])
+      apply (intro derivative_eq_intros | simp)+
+      apply (cut_tac s [of "(v - u) * x + u"])
+      using uv mult_left_le [of x "v-u"]
+      apply (auto simp:  vector_derivative_works)
+      done
+  } note vd = this
+  show ?thesis
+    apply (cut_tac has_integral_cmul [OF *, where c = "v-u"])
+    using fs assms
+    apply (simp add: False subpath_def has_contour_integral)
+    apply (rule_tac s = "(\<lambda>t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite)
+    apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real)
+    done
+qed
+
+lemma contour_integrable_subpath:
+  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
+    shows "f contour_integrable_on (subpath u v g)"
+  apply (cases u v rule: linorder_class.le_cases)
+   apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms])
+  apply (subst reversepath_subpath [symmetric])
+  apply (rule contour_integrable_reversepath)
+   using assms apply (blast intro: valid_path_subpath)
+  apply (simp add: contour_integrable_on_def)
+  using assms apply (blast intro: has_contour_integral_subpath)
+  done
+
+lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i"
+  by blast
+
+lemma has_integral_contour_integral_subpath:
+  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
+    shows "(((\<lambda>x. f(g x) * vector_derivative g (at x)))
+            has_integral  contour_integral (subpath u v g) f) {u..v}"
+  using assms
+  apply (auto simp: has_integral_integrable_integral)
+  apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified])
+  apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on)
+  done
+
+lemma contour_integral_subcontour_integral:
+  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
+    shows "contour_integral (subpath u v g) f =
+           integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))"
+  using assms has_contour_integral_subpath contour_integral_unique by blast
+
+lemma contour_integral_subpath_combine_less:
+  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
+          "u<v" "v<w"
+    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
+           contour_integral (subpath u w g) f"
+  using assms apply (auto simp: contour_integral_subcontour_integral)
+  apply (rule integral_combine, auto)
+  apply (rule integrable_on_subcbox [where a=u and b=w and s = "{0..1}", simplified])
+  apply (auto simp: contour_integrable_on)
+  done
+
+lemma contour_integral_subpath_combine:
+  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
+    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
+           contour_integral (subpath u w g) f"
+proof (cases "u\<noteq>v \<and> v\<noteq>w \<and> u\<noteq>w")
+  case True
+    have *: "subpath v u g = reversepath(subpath u v g) \<and>
+             subpath w u g = reversepath(subpath u w g) \<and>
+             subpath w v g = reversepath(subpath v w g)"
+      by (auto simp: reversepath_subpath)
+    have "u < v \<and> v < w \<or>
+          u < w \<and> w < v \<or>
+          v < u \<and> u < w \<or>
+          v < w \<and> w < u \<or>
+          w < u \<and> u < v \<or>
+          w < v \<and> v < u"
+      using True assms by linarith
+    with assms show ?thesis
+      using contour_integral_subpath_combine_less [of f g u v w]
+            contour_integral_subpath_combine_less [of f g u w v]
+            contour_integral_subpath_combine_less [of f g v u w]
+            contour_integral_subpath_combine_less [of f g v w u]
+            contour_integral_subpath_combine_less [of f g w u v]
+            contour_integral_subpath_combine_less [of f g w v u]
+      apply simp
+      apply (elim disjE)
+      apply (auto simp: * contour_integral_reversepath contour_integrable_subpath
+                   valid_path_reversepath valid_path_subpath algebra_simps)
+      done
+next
+  case False
+  then show ?thesis
+    apply (auto simp: contour_integral_subpath_refl)
+    using assms
+    by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath)
+qed
+
+lemma contour_integral_integral:
+     "contour_integral g f = integral {0..1} (\<lambda>x. f (g x) * vector_derivative g (at x))"
+  by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on)
+
+
+text\<open>Cauchy's theorem where there's a primitive\<close>
+
+lemma contour_integral_primitive_lemma:
+  fixes f :: "complex \<Rightarrow> complex" and g :: "real \<Rightarrow> complex"
+  assumes "a \<le> b"
+      and "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
+      and "g piecewise_differentiable_on {a..b}"  "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
+    shows "((\<lambda>x. f'(g x) * vector_derivative g (at x within {a..b}))
+             has_integral (f(g b) - f(g a))) {a..b}"
+proof -
+  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"
+    using assms by (auto simp: piecewise_differentiable_on_def)
+  have cfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
+    apply (rule continuous_on_compose [OF cg, unfolded o_def])
+    using assms
+    apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff)
+    done
+  { fix x::real
+    assume a: "a < x" and b: "x < b" and xk: "x \<notin> k"
+    then have "g differentiable at x within {a..b}"
+      using k by (simp add: differentiable_at_withinI)
+    then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})"
+      by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real)
+    then have gdiff: "(g has_derivative (\<lambda>u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})"
+      by (simp add: has_vector_derivative_def scaleR_conv_of_real)
+    have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})"
+      using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def)
+    then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})"
+      by (simp add: has_field_derivative_def)
+    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})"
+      using diff_chain_within [OF gdiff fdiff]
+      by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac)
+  } note * = this
+  show ?thesis
+    apply (rule fundamental_theorem_of_calculus_interior_strong)
+    using k assms cfg *
+    apply (auto simp: at_within_closed_interval)
+    done
+qed
+
+lemma contour_integral_primitive:
+  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
+      and "valid_path g" "path_image g \<subseteq> s"
+    shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g"
+  using assms
+  apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def)
+  apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s])
+  done
+
+corollary Cauchy_theorem_primitive:
+  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
+      and "valid_path g"  "path_image g \<subseteq> s" "pathfinish g = pathstart g"
+    shows "(f' has_contour_integral 0) g"
+  using assms
+  by (metis diff_self contour_integral_primitive)
+
+
+text\<open>Existence of path integral for continuous function\<close>
+lemma contour_integrable_continuous_linepath:
+  assumes "continuous_on (closed_segment a b) f"
+  shows "f contour_integrable_on (linepath a b)"
+proof -
+  have "continuous_on {0..1} ((\<lambda>x. f x * (b - a)) o linepath a b)"
+    apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01)
+    apply (rule continuous_intros | simp add: assms)+
+    done
+  then show ?thesis
+    apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric])
+    apply (rule integrable_continuous [of 0 "1::real", simplified])
+    apply (rule continuous_on_eq [where f = "\<lambda>x. f(linepath a b x)*(b - a)"])
+    apply (auto simp: vector_derivative_linepath_within)
+    done
+qed
+
+lemma has_field_der_id: "((\<lambda>x. x\<^sup>2 / 2) has_field_derivative x) (at x)"
+  by (rule has_derivative_imp_has_field_derivative)
+     (rule derivative_intros | simp)+
+
+lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\<lambda>y. y) = (b^2 - a^2)/2"
+  apply (rule contour_integral_unique)
+  using contour_integral_primitive [of UNIV "\<lambda>x. x^2/2" "\<lambda>x. x" "linepath a b"]
+  apply (auto simp: field_simps has_field_der_id)
+  done
+
+lemma contour_integrable_on_const [iff]: "(\<lambda>x. c) contour_integrable_on (linepath a b)"
+  by (simp add: continuous_on_const contour_integrable_continuous_linepath)
+
+lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)"
+  by (simp add: continuous_on_id contour_integrable_continuous_linepath)
+
+
+subsection\<open>Arithmetical combining theorems\<close>
+
+lemma has_contour_integral_neg:
+    "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g"
+  by (simp add: has_integral_neg has_contour_integral_def)
+
+lemma has_contour_integral_add:
+    "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
+     \<Longrightarrow> ((\<lambda>x. f1 x + f2 x) has_contour_integral (i1 + i2)) g"
+  by (simp add: has_integral_add has_contour_integral_def algebra_simps)
+
+lemma has_contour_integral_diff:
+  "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
+         \<Longrightarrow> ((\<lambda>x. f1 x - f2 x) has_contour_integral (i1 - i2)) g"
+  by (simp add: has_integral_sub has_contour_integral_def algebra_simps)
+
+lemma has_contour_integral_lmul:
+  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. c * (f x)) has_contour_integral (c*i)) g"
+apply (simp add: has_contour_integral_def)
+apply (drule has_integral_mult_right)
+apply (simp add: algebra_simps)
+done
+
+lemma has_contour_integral_rmul:
+  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. (f x) * c) has_contour_integral (i*c)) g"
+apply (drule has_contour_integral_lmul)
+apply (simp add: mult.commute)
+done
+
+lemma has_contour_integral_div:
+  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. f x/c) has_contour_integral (i/c)) g"
+  by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul)
+
+lemma has_contour_integral_eq:
+    "\<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"
+apply (simp add: path_image_def has_contour_integral_def)
+by (metis (no_types, lifting) image_eqI has_integral_eq)
+
+lemma has_contour_integral_bound_linepath:
+  assumes "(f has_contour_integral i) (linepath a b)"
+          "0 \<le> B" "\<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B"
+    shows "norm i \<le> B * norm(b - a)"
+proof -
+  { fix x::real
+    assume x: "0 \<le> x" "x \<le> 1"
+  have "norm (f (linepath a b x)) *
+        norm (vector_derivative (linepath a b) (at x within {0..1})) \<le> B * norm (b - a)"
+    by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x)
+  } note * = this
+  have "norm i \<le> (B * norm (b - a)) * content (cbox 0 (1::real))"
+    apply (rule has_integral_bound
+       [of _ "\<lambda>x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"])
+    using assms * unfolding has_contour_integral_def
+    apply (auto simp: norm_mult)
+    done
+  then show ?thesis
+    by (auto simp: content_real)
+qed
+
+(*UNUSED
+lemma has_contour_integral_bound_linepath_strong:
+  fixes a :: real and f :: "complex \<Rightarrow> real"
+  assumes "(f has_contour_integral i) (linepath a b)"
+          "finite k"
+          "0 \<le> B" "\<And>x::real. x \<in> closed_segment a b - k \<Longrightarrow> norm(f x) \<le> B"
+    shows "norm i \<le> B*norm(b - a)"
+*)
+
+lemma has_contour_integral_const_linepath: "((\<lambda>x. c) has_contour_integral c*(b - a))(linepath a b)"
+  unfolding has_contour_integral_linepath
+  by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one)
+
+lemma has_contour_integral_0: "((\<lambda>x. 0) has_contour_integral 0) g"
+  by (simp add: has_contour_integral_def)
+
+lemma has_contour_integral_is_0:
+    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g"
+  by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto
+
+lemma has_contour_integral_setsum:
+    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk>
+     \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p"
+  by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
+
+
+subsection \<open>Operations on path integrals\<close>
+
+lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)"
+  by (rule contour_integral_unique [OF has_contour_integral_const_linepath])
+
+lemma contour_integral_neg:
+    "f contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. -(f x)) = -(contour_integral g f)"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg)
+
+lemma contour_integral_add:
+    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x + f2 x) =
+                contour_integral g f1 + contour_integral g f2"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add)
+
+lemma contour_integral_diff:
+    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x - f2 x) =
+                contour_integral g f1 - contour_integral g f2"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff)
+
+lemma contour_integral_lmul:
+  shows "f contour_integrable_on g
+           \<Longrightarrow> contour_integral g (\<lambda>x. c * f x) = c*contour_integral g f"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul)
+
+lemma contour_integral_rmul:
+  shows "f contour_integrable_on g
+        \<Longrightarrow> contour_integral g (\<lambda>x. f x * c) = contour_integral g f * c"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul)
+
+lemma contour_integral_div:
+  shows "f contour_integrable_on g
+        \<Longrightarrow> contour_integral g (\<lambda>x. f x / c) = contour_integral g f / c"
+  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div)
+
+lemma contour_integral_eq:
+    "(\<And>x. x \<in> path_image p \<Longrightarrow> f x = g x) \<Longrightarrow> contour_integral p f = contour_integral p g"
+  apply (simp add: contour_integral_def)
+  using has_contour_integral_eq
+  by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral)
+
+lemma contour_integral_eq_0:
+    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> contour_integral g f = 0"
+  by (simp add: has_contour_integral_is_0 contour_integral_unique)
+
+lemma contour_integral_bound_linepath:
+  shows
+    "\<lbrakk>f contour_integrable_on (linepath a b);
+      0 \<le> B; \<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
+     \<Longrightarrow> norm(contour_integral (linepath a b) f) \<le> B*norm(b - a)"
+  apply (rule has_contour_integral_bound_linepath [of f])
+  apply (auto simp: has_contour_integral_integral)
+  done
+
+lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0"
+  by (simp add: contour_integral_unique has_contour_integral_0)
+
+lemma contour_integral_setsum:
+    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
+     \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s"
+  by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral)
+
+lemma contour_integrable_eq:
+    "\<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"
+  unfolding contour_integrable_on_def
+  by (metis has_contour_integral_eq)
+
+
+subsection \<open>Arithmetic theorems for path integrability\<close>
+
+lemma contour_integrable_neg:
+    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g"
+  using has_contour_integral_neg contour_integrable_on_def by blast
+
+lemma contour_integrable_add:
+    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x + f2 x) contour_integrable_on g"
+  using has_contour_integral_add contour_integrable_on_def
+  by fastforce
+
+lemma contour_integrable_diff:
+    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x - f2 x) contour_integrable_on g"
+  using has_contour_integral_diff contour_integrable_on_def
+  by fastforce
+
+lemma contour_integrable_lmul:
+    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. c * f x) contour_integrable_on g"
+  using has_contour_integral_lmul contour_integrable_on_def
+  by fastforce
+
+lemma contour_integrable_rmul:
+    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x * c) contour_integrable_on g"
+  using has_contour_integral_rmul contour_integrable_on_def
+  by fastforce
+
+lemma contour_integrable_div:
+    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x / c) contour_integrable_on g"
+  using has_contour_integral_div contour_integrable_on_def
+  by fastforce
+
+lemma contour_integrable_setsum:
+    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
+     \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p"
+   unfolding contour_integrable_on_def
+   by (metis has_contour_integral_setsum)
+
+
+subsection\<open>Reversing a path integral\<close>
+
+lemma has_contour_integral_reverse_linepath:
+    "(f has_contour_integral i) (linepath a b)
+     \<Longrightarrow> (f has_contour_integral (-i)) (linepath b a)"
+  using has_contour_integral_reversepath valid_path_linepath by fastforce
+
+lemma contour_integral_reverse_linepath:
+    "continuous_on (closed_segment a b) f
+     \<Longrightarrow> contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)"
+apply (rule contour_integral_unique)
+apply (rule has_contour_integral_reverse_linepath)
+by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral)
+
+
+(* Splitting a path integral in a flat way.*)
+
+lemma has_contour_integral_split:
+  assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)"
+      and k: "0 \<le> k" "k \<le> 1"
+      and c: "c - a = k *\<^sub>R (b - a)"
+    shows "(f has_contour_integral (i + j)) (linepath a b)"
+proof (cases "k = 0 \<or> k = 1")
+  case True
+  then show ?thesis
+    using assms
+    apply auto
+    apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique)
+    apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique)
+    done
+next
+  case False
+  then have k: "0 < k" "k < 1" "complex_of_real k \<noteq> 1"
+    using assms apply auto
+    using of_real_eq_iff by fastforce
+  have c': "c = k *\<^sub>R (b - a) + a"
+    by (metis diff_add_cancel c)
+  have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)"
+    by (simp add: algebra_simps c')
+  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}"
+    have **: "\<And>x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b"
+      using False
+      apply (simp add: c' algebra_simps)
+      apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps)
+      done
+    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}"
+      using * k
+      apply -
+      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified])
+      apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c)
+      apply (drule has_integral_cmul [where c = "inverse k"])
+      apply (simp add: has_integral_cmul)
+      done
+  } note fi = this
+  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}"
+    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)"
+      using k
+      apply (simp add: c' field_simps)
+      apply (simp add: scaleR_conv_of_real divide_simps)
+      apply (simp add: field_simps)
+      done
+    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}"
+      using * k
+      apply -
+      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified])
+      apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc)
+      apply (drule has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"])
+      apply (simp add: has_integral_cmul)
+      done
+  } note fj = this
+  show ?thesis
+    using f k
+    apply (simp add: has_contour_integral_linepath)
+    apply (simp add: linepath_def)
+    apply (rule has_integral_combine [OF _ _ fi fj], simp_all)
+    done
+qed
+
+lemma continuous_on_closed_segment_transform:
+  assumes f: "continuous_on (closed_segment a b) f"
+      and k: "0 \<le> k" "k \<le> 1"
+      and c: "c - a = k *\<^sub>R (b - a)"
+    shows "continuous_on (closed_segment a c) f"
+proof -
+  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
+    using c by (simp add: algebra_simps)
+  show "continuous_on (closed_segment a c) f"
+    apply (rule continuous_on_subset [OF f])
+    apply (simp add: segment_convex_hull)
+    apply (rule convex_hull_subset)
+    using assms
+    apply (auto simp: hull_inc c' Convex.convexD_alt)
+    done
+qed
+
+lemma contour_integral_split:
+  assumes f: "continuous_on (closed_segment a b) f"
+      and k: "0 \<le> k" "k \<le> 1"
+      and c: "c - a = k *\<^sub>R (b - a)"
+    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
+proof -
+  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
+    using c by (simp add: algebra_simps)
+  have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f"
+    apply (rule_tac [!] continuous_on_subset [OF f])
+    apply (simp_all add: segment_convex_hull)
+    apply (rule_tac [!] convex_hull_subset)
+    using assms
+    apply (auto simp: hull_inc c' Convex.convexD_alt)
+    done
+  show ?thesis
+    apply (rule contour_integral_unique)
+    apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c])
+    apply (rule contour_integrable_continuous_linepath *)+
+    done
+qed
+
+lemma contour_integral_split_linepath:
+  assumes f: "continuous_on (closed_segment a b) f"
+      and c: "c \<in> closed_segment a b"
+    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
+  using c
+  by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f])
+
+(* The special case of midpoints used in the main quadrisection.*)
+
+lemma has_contour_integral_midpoint:
+  assumes "(f has_contour_integral i) (linepath a (midpoint a b))"
+          "(f has_contour_integral j) (linepath (midpoint a b) b)"
+    shows "(f has_contour_integral (i + j)) (linepath a b)"
+  apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"])
+  using assms
+  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
+  done
+
+lemma contour_integral_midpoint:
+   "continuous_on (closed_segment a b) f
+    \<Longrightarrow> contour_integral (linepath a b) f =
+        contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f"
+  apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"])
+  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
+  done
+
+
+text\<open>A couple of special case lemmas that are useful below\<close>
+
+lemma triangle_linear_has_chain_integral:
+    "((\<lambda>x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+  apply (rule Cauchy_theorem_primitive [of UNIV "\<lambda>x. m/2 * x^2 + d*x"])
+  apply (auto intro!: derivative_eq_intros)
+  done
+
+lemma has_chain_integral_chain_integral3:
+     "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d)
+      \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i"
+  apply (subst contour_integral_unique [symmetric], assumption)
+  apply (drule has_contour_integral_integrable)
+  apply (simp add: valid_path_join)
+  done
+
+lemma has_chain_integral_chain_integral4:
+     "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e)
+      \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i"
+  apply (subst contour_integral_unique [symmetric], assumption)
+  apply (drule has_contour_integral_integrable)
+  apply (simp add: valid_path_join)
+  done
+
+subsection\<open>Reversing the order in a double path integral\<close>
+
+text\<open>The condition is stronger than needed but it's often true in typical situations\<close>
+
+lemma fst_im_cbox [simp]: "cbox c d \<noteq> {} \<Longrightarrow> (fst ` cbox (a,c) (b,d)) = cbox a b"
+  by (auto simp: cbox_Pair_eq)
+
+lemma snd_im_cbox [simp]: "cbox a b \<noteq> {} \<Longrightarrow> (snd ` cbox (a,c) (b,d)) = cbox c d"
+  by (auto simp: cbox_Pair_eq)
+
+lemma contour_integral_swap:
+  assumes fcon:  "continuous_on (path_image g \<times> path_image h) (\<lambda>(y1,y2). f y1 y2)"
+      and vp:    "valid_path g" "valid_path h"
+      and gvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative g (at t))"
+      and hvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative h (at t))"
+  shows "contour_integral g (\<lambda>w. contour_integral h (f w)) =
+         contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
+proof -
+  have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h"
+    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
+  have fgh1: "\<And>x. (\<lambda>t. f (g x) (h t)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g x, h t))"
+    by (rule ext) simp
+  have fgh2: "\<And>x. (\<lambda>t. f (g t) (h x)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g t, h x))"
+    by (rule ext) simp
+  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)"
+    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
+  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)"
+    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
+  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}"
+    apply (rule integrable_continuous_real)
+    apply (rule continuous_on_mult [OF _ gvcon])
+    apply (subst fgh2)
+    apply (rule fcon_im2 gcon continuous_intros | simp)+
+    done
+  have "(\<lambda>z. vector_derivative g (at (fst z))) = (\<lambda>x. vector_derivative g (at x)) o fst"
+    by auto
+  then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>x. vector_derivative g (at (fst x)))"
+    apply (rule ssubst)
+    apply (rule continuous_intros | simp add: gvcon)+
+    done
+  have "(\<lambda>z. vector_derivative h (at (snd z))) = (\<lambda>x. vector_derivative h (at x)) o snd"
+    by auto
+  then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))"
+    apply (rule ssubst)
+    apply (rule continuous_intros | simp add: hvcon)+
+    done
+  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))"
+    by auto
+  then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))"
+    apply (rule ssubst)
+    apply (rule gcon hcon continuous_intros | simp)+
+    apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
+    done
+  have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) =
+        integral {0..1} (\<lambda>x. contour_integral h (\<lambda>y. f (g x) y * vector_derivative g (at x)))"
+    apply (rule integral_cong [OF contour_integral_rmul [symmetric]])
+    apply (clarsimp simp: contour_integrable_on)
+    apply (rule integrable_continuous_real)
+    apply (rule continuous_on_mult [OF _ hvcon])
+    apply (subst fgh1)
+    apply (rule fcon_im1 hcon continuous_intros | simp)+
+    done
+  also have "... = integral {0..1}
+                     (\<lambda>y. contour_integral g (\<lambda>x. f x (h y) * vector_derivative h (at y)))"
+    apply (simp only: contour_integral_integral)
+    apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified])
+     apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+
+    unfolding integral_mult_left [symmetric]
+    apply (simp only: mult_ac)
+    done
+  also have "... = contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
+    apply (simp add: contour_integral_integral)
+    apply (rule integral_cong)
+    unfolding integral_mult_left [symmetric]
+    apply (simp add: algebra_simps)
+    done
+  finally show ?thesis
+    by (simp add: contour_integral_integral)
+qed
+
+
+subsection\<open>The key quadrisection step\<close>
+
+lemma norm_sum_half:
+  assumes "norm(a + b) >= e"
+    shows "norm a >= e/2 \<or> norm b >= e/2"
+proof -
+  have "e \<le> norm (- a - b)"
+    by (simp add: add.commute assms norm_minus_commute)
+  thus ?thesis
+    using norm_triangle_ineq4 order_trans by fastforce
+qed
+
+lemma norm_sum_lemma:
+  assumes "e \<le> norm (a + b + c + d)"
+    shows "e / 4 \<le> norm a \<or> e / 4 \<le> norm b \<or> e / 4 \<le> norm c \<or> e / 4 \<le> norm d"
+proof -
+  have "e \<le> norm ((a + b) + (c + d))" using assms
+    by (simp add: algebra_simps)
+  then show ?thesis
+    by (auto dest!: norm_sum_half)
+qed
+
+lemma Cauchy_theorem_quadrisection:
+  assumes f: "continuous_on (convex hull {a,b,c}) f"
+      and dist: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
+      and e: "e * K^2 \<le>
+              norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)"
+  shows "\<exists>a' b' c'.
+           a' \<in> convex hull {a,b,c} \<and> b' \<in> convex hull {a,b,c} \<and> c' \<in> convex hull {a,b,c} \<and>
+           dist a' b' \<le> K/2  \<and>  dist b' c' \<le> K/2  \<and>  dist c' a' \<le> K/2  \<and>
+           e * (K/2)^2 \<le> norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)"
+proof -
+  note divide_le_eq_numeral1 [simp del]
+  define a' where "a' = midpoint b c"
+  define b' where "b' = midpoint c a"
+  define c' where "c' = midpoint a b"
+  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
+    using f continuous_on_subset segments_subset_convex_hull by metis+
+  have fcont': "continuous_on (closed_segment c' b') f"
+               "continuous_on (closed_segment a' c') f"
+               "continuous_on (closed_segment b' a') f"
+    unfolding a'_def b'_def c'_def
+    apply (rule continuous_on_subset [OF f],
+           metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+
+    done
+  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
+  have *: "?pathint a b + ?pathint b c + ?pathint c a =
+          (?pathint a c' + ?pathint c' b' + ?pathint b' a) +
+          (?pathint a' c' + ?pathint c' b + ?pathint b a') +
+          (?pathint a' c + ?pathint c b' + ?pathint b' a') +
+          (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
+    apply (simp add: fcont' contour_integral_reverse_linepath)
+    apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc)
+    done
+  have [simp]: "\<And>x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2"
+    by (metis left_diff_distrib mult.commute norm_mult_numeral1)
+  have [simp]: "\<And>x y. cmod (x - y) = cmod (y - x)"
+    by (simp add: norm_minus_commute)
+  consider "e * K\<^sup>2 / 4 \<le> cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" |
+           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" |
+           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" |
+           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
+    using assms
+    apply (simp only: *)
+    apply (blast intro: that dest!: norm_sum_lemma)
+    done
+  then show ?thesis
+  proof cases
+    case 1 then show ?thesis
+      apply (rule_tac x=a in exI)
+      apply (rule exI [where x=c'])
+      apply (rule exI [where x=b'])
+      using assms
+      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
+      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
+      done
+  next
+    case 2 then show ?thesis
+      apply (rule_tac x=a' in exI)
+      apply (rule exI [where x=c'])
+      apply (rule exI [where x=b])
+      using assms
+      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
+      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
+      done
+  next
+    case 3 then show ?thesis
+      apply (rule_tac x=a' in exI)
+      apply (rule exI [where x=c])
+      apply (rule exI [where x=b'])
+      using assms
+      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
+      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
+      done
+  next
+    case 4 then show ?thesis
+      apply (rule_tac x=a' in exI)
+      apply (rule exI [where x=b'])
+      apply (rule exI [where x=c'])
+      using assms
+      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
+      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
+      done
+  qed
+qed
+
+subsection\<open>Cauchy's theorem for triangles\<close>
+
+lemma triangle_points_closer:
+  fixes a::complex
+  shows "\<lbrakk>x \<in> convex hull {a,b,c};  y \<in> convex hull {a,b,c}\<rbrakk>
+         \<Longrightarrow> norm(x - y) \<le> norm(a - b) \<or>
+             norm(x - y) \<le> norm(b - c) \<or>
+             norm(x - y) \<le> norm(c - a)"
+  using simplex_extremal_le [of "{a,b,c}"]
+  by (auto simp: norm_minus_commute)
+
+lemma holomorphic_point_small_triangle:
+  assumes x: "x \<in> s"
+      and f: "continuous_on s f"
+      and cd: "f field_differentiable (at x within s)"
+      and e: "0 < e"
+    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>
+              x \<in> convex hull {a,b,c} \<and> convex hull {a,b,c} \<subseteq> s
+              \<longrightarrow> norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f +
+                       contour_integral(linepath c a) f)
+                  \<le> e*(dist a b + dist b c + dist c a)^2"
+           (is "\<exists>k>0. \<forall>a b c. _ \<longrightarrow> ?normle a b c")
+proof -
+  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>
+                     \<Longrightarrow> a \<le> e*(x + y + z)^2"
+    by (simp add: algebra_simps power2_eq_square)
+  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"
+             for x::real and a b c
+    by linarith
+  have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a"
+              if "convex hull {a, b, c} \<subseteq> s" for a b c
+    using segments_subset_convex_hull that
+    by (metis continuous_on_subset f contour_integrable_continuous_linepath)+
+  note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral]
+  { fix f' a b c d
+    assume d: "0 < d"
+       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)"
+       and le: "cmod (a - b) \<le> d" "cmod (b - c) \<le> d" "cmod (c - a) \<le> d"
+       and xc: "x \<in> convex hull {a, b, c}"
+       and s: "convex hull {a, b, c} \<subseteq> s"
+    have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f =
+              contour_integral (linepath a b) (\<lambda>y. f y - f x - f'*(y - x)) +
+              contour_integral (linepath b c) (\<lambda>y. f y - f x - f'*(y - x)) +
+              contour_integral (linepath c a) (\<lambda>y. f y - f x - f'*(y - x))"
+      apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s])
+      apply (simp add: field_simps)
+      done
+    { fix y
+      assume yc: "y \<in> convex hull {a,b,c}"
+      have "cmod (f y - f x - f' * (y - x)) \<le> e*norm(y - x)"
+        apply (rule f')
+        apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans)
+        using s yc by blast
+      also have "... \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))"
+        by (simp add: yc e xc disj_le [OF triangle_points_closer])
+      finally have "cmod (f y - f x - f' * (y - x)) \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" .
+    } note cm_le = this
+    have "?normle a b c"
+      apply (simp add: dist_norm pa)
+      apply (rule le_of_3)
+      using f' xc s e
+      apply simp_all
+      apply (intro norm_triangle_le add_mono path_bound)
+      apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc)
+      apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+
+      done
+  } note * = this
+  show ?thesis
+    using cd e
+    apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def)
+    apply (clarify dest!: spec mp)
+    using *
+    apply (simp add: dist_norm, blast)
+    done
+qed
+
+
+(* Hence the most basic theorem for a triangle.*)
+locale Chain =
+  fixes x0 At Follows
+  assumes At0: "At x0 0"
+      and AtSuc: "\<And>x n. At x n \<Longrightarrow> \<exists>x'. At x' (Suc n) \<and> Follows x' x"
+begin
+  primrec f where
+    "f 0 = x0"
+  | "f (Suc n) = (SOME x. At x (Suc n) \<and> Follows x (f n))"
+
+  lemma At: "At (f n) n"
+  proof (induct n)
+    case 0 show ?case
+      by (simp add: At0)
+  next
+    case (Suc n) show ?case
+      by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex)
+  qed
+
+  lemma Follows: "Follows (f(Suc n)) (f n)"
+    by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex)
+
+  declare f.simps(2) [simp del]
+end
+
+lemma Chain3:
+  assumes At0: "At x0 y0 z0 0"
+      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"
+  obtains f g h where
+    "f 0 = x0" "g 0 = y0" "h 0 = z0"
+                      "\<And>n. At (f n) (g n) (h n) n"
+                       "\<And>n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)"
+proof -
+  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"
+    apply unfold_locales
+    using At0 AtSuc by auto
+  show ?thesis
+  apply (rule that [of "\<lambda>n. fst (three.f n)"  "\<lambda>n. fst (snd (three.f n))" "\<lambda>n. snd (snd (three.f n))"])
+  apply simp_all
+  using three.At three.Follows
+  apply (simp_all add: split_beta')
+  done
+qed
+
+lemma Cauchy_theorem_triangle:
+  assumes "f holomorphic_on (convex hull {a,b,c})"
+    shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+proof -
+  have contf: "continuous_on (convex hull {a,b,c}) f"
+    by (metis assms holomorphic_on_imp_continuous_on)
+  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
+  { fix y::complex
+    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
+       and ynz: "y \<noteq> 0"
+    define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))"
+    define e where "e = norm y / K^2"
+    have K1: "K \<ge> 1"  by (simp add: K_def max.coboundedI1)
+    then have K: "K > 0" by linarith
+    have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
+      by (simp_all add: K_def)
+    have e: "e > 0"
+      unfolding e_def using ynz K1 by simp
+    define At where "At x y z n \<longleftrightarrow>
+        convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
+        dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
+        norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2"
+      for x y z n
+    have At0: "At a b c 0"
+      using fy
+      by (simp add: At_def e_def has_chain_integral_chain_integral3)
+    { fix x y z n
+      assume At: "At x y z n"
+      then have contf': "continuous_on (convex hull {x,y,z}) f"
+        using contf At_def continuous_on_subset by blast
+      have "\<exists>x' y' z'. At x' y' z' (Suc n) \<and> convex hull {x',y',z'} \<subseteq> convex hull {x,y,z}"
+        using At
+        apply (simp add: At_def)
+        using  Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e]
+        apply clarsimp
+        apply (rule_tac x="a'" in exI)
+        apply (rule_tac x="b'" in exI)
+        apply (rule_tac x="c'" in exI)
+        apply (simp add: algebra_simps)
+        apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE)
+        done
+    } note AtSuc = this
+    obtain fa fb fc
+      where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c"
+        and cosb: "\<And>n. convex hull {fa n, fb n, fc n} \<subseteq> convex hull {a,b,c}"
+        and dist: "\<And>n. dist (fa n) (fb n) \<le> K/2^n"
+                  "\<And>n. dist (fb n) (fc n) \<le> K/2^n"
+                  "\<And>n. dist (fc n) (fa n) \<le> K/2^n"
+        and no: "\<And>n. norm(?pathint (fa n) (fb n) +
+                           ?pathint (fb n) (fc n) +
+                           ?pathint (fc n) (fa n)) \<ge> e * (K/2^n)^2"
+        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}"
+      apply (rule Chain3 [of At, OF At0 AtSuc])
+      apply (auto simp: At_def)
+      done
+    have "\<exists>x. \<forall>n. x \<in> convex hull {fa n, fb n, fc n}"
+      apply (rule bounded_closed_nest)
+      apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull)
+      apply (rule allI)
+      apply (rule transitive_stepwise_le)
+      apply (auto simp: conv_le)
+      done
+    then obtain x where x: "\<And>n. x \<in> convex hull {fa n, fb n, fc n}" by auto
+    then have xin: "x \<in> convex hull {a,b,c}"
+      using assms f0 by blast
+    then have fx: "f field_differentiable at x within (convex hull {a,b,c})"
+      using assms holomorphic_on_def by blast
+    { fix k n
+      assume k: "0 < k"
+         and le:
+            "\<And>x' y' z'.
+               \<lbrakk>dist x' y' \<le> k; dist y' z' \<le> k; dist z' x' \<le> k;
+                x \<in> convex hull {x',y',z'};
+                convex hull {x',y',z'} \<subseteq> convex hull {a,b,c}\<rbrakk>
+               \<Longrightarrow>
+               cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10
+                     \<le> e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2"
+         and Kk: "K / k < 2 ^ n"
+      have "K / 2 ^ n < k" using Kk k
+        by (auto simp: field_simps)
+      then have DD: "dist (fa n) (fb n) \<le> k" "dist (fb n) (fc n) \<le> k" "dist (fc n) (fa n) \<le> k"
+        using dist [of n]  k
+        by linarith+
+      have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2
+               \<le> (3 * K / 2 ^ n)\<^sup>2"
+        using dist [of n] e K
+        by (simp add: abs_le_square_iff [symmetric])
+      have less10: "\<And>x y::real. 0 < x \<Longrightarrow> y \<le> 9*x \<Longrightarrow> y < x*10"
+        by linarith
+      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"
+        using ynz dle e mult_le_cancel_left_pos by blast
+      also have "... <
+          cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10"
+        using no [of n] e K
+        apply (simp add: e_def field_simps)
+        apply (simp only: zero_less_norm_iff [symmetric])
+        done
+      finally have False
+        using le [OF DD x cosb] by auto
+    } then
+    have ?thesis
+      using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e
+      apply clarsimp
+      apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]])
+      apply force+
+      done
+  }
+  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
+    by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1)
+                   segments_subset_convex_hull(3) segments_subset_convex_hull(5))
+  ultimately show ?thesis
+    using has_contour_integral_integral by fastforce
+qed
+
+
+subsection\<open>Version needing function holomorphic in interior only\<close>
+
+lemma Cauchy_theorem_flat_lemma:
+  assumes f: "continuous_on (convex hull {a,b,c}) f"
+      and c: "c - a = k *\<^sub>R (b - a)"
+      and k: "0 \<le> k"
+    shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f +
+          contour_integral (linepath c a) f = 0"
+proof -
+  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
+    using f continuous_on_subset segments_subset_convex_hull by metis+
+  show ?thesis
+  proof (cases "k \<le> 1")
+    case True show ?thesis
+      by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc)
+  next
+    case False then show ?thesis
+      using fabc c
+      apply (subst contour_integral_split [of a c f "1/k" b, symmetric])
+      apply (metis closed_segment_commute fabc(3))
+      apply (auto simp: k contour_integral_reverse_linepath)
+      done
+  qed
+qed
+
+lemma Cauchy_theorem_flat:
+  assumes f: "continuous_on (convex hull {a,b,c}) f"
+      and c: "c - a = k *\<^sub>R (b - a)"
+    shows "contour_integral (linepath a b) f +
+           contour_integral (linepath b c) f +
+           contour_integral (linepath c a) f = 0"
+proof (cases "0 \<le> k")
+  case True with assms show ?thesis
+    by (blast intro: Cauchy_theorem_flat_lemma)
+next
+  case False
+  have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
+    using f continuous_on_subset segments_subset_convex_hull by metis+
+  moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f +
+        contour_integral (linepath c b) f = 0"
+    apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"])
+    using False c
+    apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps)
+    done
+  ultimately show ?thesis
+    apply (auto simp: contour_integral_reverse_linepath)
+    using add_eq_0_iff by force
+qed
+
+
+lemma Cauchy_theorem_triangle_interior:
+  assumes contf: "continuous_on (convex hull {a,b,c}) f"
+      and holf:  "f holomorphic_on interior (convex hull {a,b,c})"
+     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+proof -
+  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
+    using contf continuous_on_subset segments_subset_convex_hull by metis+
+  have "bounded (f ` (convex hull {a,b,c}))"
+    by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf)
+  then obtain B where "0 < B" and Bnf: "\<And>x. x \<in> convex hull {a,b,c} \<Longrightarrow> norm (f x) \<le> B"
+     by (auto simp: dest!: bounded_pos [THEN iffD1])
+  have "bounded (convex hull {a,b,c})"
+    by (simp add: bounded_convex_hull)
+  then obtain C where C: "0 < C" and Cno: "\<And>y. y \<in> convex hull {a,b,c} \<Longrightarrow> norm y < C"
+    using bounded_pos_less by blast
+  then have diff_2C: "norm(x - y) \<le> 2*C"
+           if x: "x \<in> convex hull {a, b, c}" and y: "y \<in> convex hull {a, b, c}" for x y
+  proof -
+    have "cmod x \<le> C"
+      using x by (meson Cno not_le not_less_iff_gr_or_eq)
+    hence "cmod (x - y) \<le> C + C"
+      using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans)
+    thus "cmod (x - y) \<le> 2 * C"
+      by (metis mult_2)
+  qed
+  have contf': "continuous_on (convex hull {b,a,c}) f"
+    using contf by (simp add: insert_commute)
+  { fix y::complex
+    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
+       and ynz: "y \<noteq> 0"
+    have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y"
+      by (rule has_chain_integral_chain_integral3 [OF fy])
+    have ?thesis
+    proof (cases "c=a \<or> a=b \<or> b=c")
+      case True then show ?thesis
+        using Cauchy_theorem_flat [OF contf, of 0]
+        using has_chain_integral_chain_integral3 [OF fy] ynz
+        by (force simp: fabc contour_integral_reverse_linepath)
+    next
+      case False
+      then have car3: "card {a, b, c} = Suc (DIM(complex))"
+        by auto
+      { assume "interior(convex hull {a,b,c}) = {}"
+        then have "collinear{a,b,c}"
+          using interior_convex_hull_eq_empty [OF car3]
+          by (simp add: collinear_3_eq_affine_dependent)
+        then have "False"
+          using False
+          apply (clarsimp simp add: collinear_3 collinear_lemma)
+          apply (drule Cauchy_theorem_flat [OF contf'])
+          using pi_eq_y ynz
+          apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath)
+          done
+      }
+      then obtain d where d: "d \<in> interior (convex hull {a, b, c})"
+        by blast
+      { fix d1
+        assume d1_pos: "0 < d1"
+           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>
+                           \<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)"
+        define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))"
+        define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x
+        let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
+        have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B"
+          using d1_pos \<open>C>0\<close> \<open>B>0\<close> ynz by (simp_all add: e_def)
+        then have eCB: "24 * e * C * B \<le> cmod y"
+          using \<open>C>0\<close> \<open>B>0\<close>  by (simp add: field_simps)
+        have e_le_d1: "e * (4 * C) \<le> d1"
+          using e \<open>C>0\<close> by (simp add: field_simps)
+        have "shrink a \<in> interior(convex hull {a,b,c})"
+             "shrink b \<in> interior(convex hull {a,b,c})"
+             "shrink c \<in> interior(convex hull {a,b,c})"
+          using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
+        then have fhp0: "(f has_contour_integral 0)
+                (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))"
+          by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior)
+        then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0"
+          by (simp add: has_chain_integral_chain_integral3)
+        have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)"
+                      "f contour_integrable_on linepath (shrink b) (shrink c)"
+                      "f contour_integrable_on linepath (shrink c) (shrink a)"
+          using fhp0  by (auto simp: valid_path_join dest: has_contour_integral_integrable)
+        have cmod_shr: "\<And>x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)"
+          using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric])
+        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)"
+          by (simp add: algebra_simps)
+        have "cmod y / (24 * C) \<le> cmod y / cmod (b - a) / 12"
+          using False \<open>C>0\<close> diff_2C [of b a] ynz
+          by (auto simp: divide_simps hull_inc)
+        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
+          apply (cases "x=0", simp add: \<open>0<C\<close>)
+          using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast
+        { fix u v
+          assume uv: "u \<in> convex hull {a, b, c}" "v \<in> convex hull {a, b, c}" "u\<noteq>v"
+             and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)"
+          have shr_uv: "shrink u \<in> interior(convex hull {a,b,c})"
+                       "shrink v \<in> interior(convex hull {a,b,c})"
+            using d e uv
+            by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
+          have cmod_fuv: "\<And>x. 0\<le>x \<Longrightarrow> x\<le>1 \<Longrightarrow> cmod (f (linepath (shrink u) (shrink v) x)) \<le> B"
+            using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD])
+          have By_uv: "B * (12 * (e * cmod (u - v))) \<le> cmod y"
+            apply (rule order_trans [OF _ eCB])
+            using e \<open>B>0\<close> diff_2C [of u v] uv
+            by (auto simp: field_simps)
+          { fix x::real   assume x: "0\<le>x" "x\<le>1"
+            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)"
+              apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0])
+              using uv x d interior_subset
+              apply (auto simp: hull_inc intro!: less_C)
+              done
+            have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))"
+              by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real)
+            have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1"
+              using \<open>e>0\<close>
+              apply (simp add: ll norm_mult scaleR_diff_right)
+              apply (rule less_le_trans [OF _ e_le_d1])
+              using cmod_less_4C
+              apply (force intro: norm_triangle_lt)
+              done
+            have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)"
+              using x uv shr_uv cmod_less_dt
+              by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull)
+            also have "... \<le> cmod y / cmod (v - u) / 12"
+              using False uv \<open>C>0\<close> diff_2C [of v u] ynz
+              by (auto simp: divide_simps hull_inc)
+            finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \<le> cmod y / cmod (v - u) / 12"
+              by simp
+            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"
+              using uv False by (auto simp: field_simps)
+            have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) +
+                  cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x))
+                  \<le> cmod y / 6"
+              apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"])
+              apply (rule add_mono [OF mult_mono])
+              using By_uv e \<open>0 < B\<close> \<open>0 < C\<close> x ynz
+              apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc)
+              apply (simp add: field_simps)
+              done
+          } note cmod_diff_le = this
+          have f_uv: "continuous_on (closed_segment u v) f"
+            by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull])
+          have **: "\<And>f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)"
+            by (simp add: algebra_simps)
+          have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \<le> norm y / 6"
+            apply (rule order_trans)
+            apply (rule has_integral_bound
+                    [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)"
+                        "\<lambda>x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)"
+                        _ 0 1 ])
+            using ynz \<open>0 < B\<close> \<open>0 < C\<close>
+            apply (simp_all del: le_divide_eq_numeral1)
+            apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral
+                             fpi_uv f_uv contour_integrable_continuous_linepath, clarify)
+            apply (simp only: **)
+            apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1)
+            done
+          } note * = this
+          have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \<le> norm y / 6"
+            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
+          moreover
+          have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \<le> norm y / 6"
+            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
+          moreover
+          have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \<le> norm y / 6"
+            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
+          ultimately
+          have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) +
+                     (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a))
+                \<le> norm y / 6 + norm y / 6 + norm y / 6"
+            by (metis norm_triangle_le add_mono)
+          also have "... = norm y / 2"
+            by simp
+          finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) -
+                          (?pathint a b + ?pathint b c + ?pathint c a))
+                \<le> norm y / 2"
+            by (simp add: algebra_simps)
+          then
+          have "norm(?pathint a b + ?pathint b c + ?pathint c a) \<le> norm y / 2"
+            by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff)
+          then have "False"
+            using pi_eq_y ynz by auto
+        }
+        moreover have "uniformly_continuous_on (convex hull {a,b,c}) f"
+          by (simp add: contf compact_convex_hull compact_uniformly_continuous)
+        ultimately have "False"
+          unfolding uniformly_continuous_on_def
+          by (force simp: ynz \<open>0 < C\<close> dist_norm)
+        then show ?thesis ..
+      qed
+  }
+  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
+    using fabc contour_integrable_continuous_linepath by auto
+  ultimately show ?thesis
+    using has_contour_integral_integral by fastforce
+qed
+
+
+subsection\<open>Version allowing finite number of exceptional points\<close>
+
+lemma Cauchy_theorem_triangle_cofinite:
+  assumes "continuous_on (convex hull {a,b,c}) f"
+      and "finite s"
+      and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - s \<Longrightarrow> f field_differentiable (at x))"
+     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+using assms
+proof (induction "card s" arbitrary: a b c s rule: less_induct)
+  case (less s a b c)
+  show ?case
+  proof (cases "s={}")
+    case True with less show ?thesis
+      by (fastforce simp: holomorphic_on_def field_differentiable_at_within
+                    Cauchy_theorem_triangle_interior)
+  next
+    case False
+    then obtain d s' where d: "s = insert d s'" "d \<notin> s'"
+      by (meson Set.set_insert all_not_in_conv)
+    then show ?thesis
+    proof (cases "d \<in> convex hull {a,b,c}")
+      case False
+      show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+        apply (rule less.hyps [of "s'"])
+        using False d \<open>finite s\<close> interior_subset
+        apply (auto intro!: less.prems)
+        done
+    next
+      case True
+      have *: "convex hull {a, b, d} \<subseteq> convex hull {a, b, c}"
+        by (meson True hull_subset insert_subset convex_hull_subset)
+      have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)"
+        apply (rule less.hyps [of "s'"])
+        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
+        apply (auto intro!: less.prems continuous_on_subset [OF  _ *])
+        apply (metis * insert_absorb insert_subset interior_mono)
+        done
+      have *: "convex hull {b, c, d} \<subseteq> convex hull {a, b, c}"
+        by (meson True hull_subset insert_subset convex_hull_subset)
+      have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)"
+        apply (rule less.hyps [of "s'"])
+        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
+        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
+        apply (metis * insert_absorb insert_subset interior_mono)
+        done
+      have *: "convex hull {c, a, d} \<subseteq> convex hull {a, b, c}"
+        by (meson True hull_subset insert_subset convex_hull_subset)
+      have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)"
+        apply (rule less.hyps [of "s'"])
+        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
+        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
+        apply (metis * insert_absorb insert_subset interior_mono)
+        done
+      have "f contour_integrable_on linepath a b"
+        using less.prems
+        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
+      moreover have "f contour_integrable_on linepath b c"
+        using less.prems
+        by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3))
+      moreover have "f contour_integrable_on linepath c a"
+        using less.prems
+        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
+      ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
+        by auto
+      { fix y::complex
+        assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
+           and ynz: "y \<noteq> 0"
+        have cont_ad: "continuous_on (closed_segment a d) f"
+          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3))
+        have cont_bd: "continuous_on (closed_segment b d) f"
+          by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1))
+        have cont_cd: "continuous_on (closed_segment c d) f"
+          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2))
+        have "contour_integral  (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))"
+                "contour_integral  (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))"
+                "contour_integral  (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)"
+            using has_chain_integral_chain_integral3 [OF abd]
+                  has_chain_integral_chain_integral3 [OF bcd]
+                  has_chain_integral_chain_integral3 [OF cad]
+            by (simp_all add: algebra_simps add_eq_0_iff)
+        then have ?thesis
+          using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce
+      }
+      then show ?thesis
+        using fpi contour_integrable_on_def by blast
+    qed
+  qed
+qed
+
+
+subsection\<open>Cauchy's theorem for an open starlike set\<close>
+
+lemma starlike_convex_subset:
+  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"
+    shows "convex hull {a,b,c} \<subseteq> s"
+      using s
+      apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull)
+      apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE)
+      done
+
+lemma triangle_contour_integrals_starlike_primitive:
+  assumes contf: "continuous_on s f"
+      and s: "a \<in> s" "open s"
+      and x: "x \<in> s"
+      and subs: "\<And>y. y \<in> s \<Longrightarrow> closed_segment a y \<subseteq> s"
+      and zer: "\<And>b c. closed_segment b c \<subseteq> s
+                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
+                       contour_integral (linepath c a) f = 0"
+    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)"
+proof -
+  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
+  { fix e y
+    assume e: "0 < e" and bxe: "ball x e \<subseteq> s" and close: "cmod (y - x) < e"
+    have y: "y \<in> s"
+      using bxe close  by (force simp: dist_norm norm_minus_commute)
+    have cont_ayf: "continuous_on (closed_segment a y) f"
+      using contf continuous_on_subset subs y by blast
+    have xys: "closed_segment x y \<subseteq> s"
+      apply (rule order_trans [OF _ bxe])
+      using close
+      by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound)
+    have "?pathint a y - ?pathint a x = ?pathint x y"
+      using zer [OF xys]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
+  } note [simp] = this
+  { fix e::real
+    assume e: "0 < e"
+    have cont_atx: "continuous (at x) f"
+      using x s contf continuous_on_eq_continuous_at by blast
+    then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. cmod (y - x) < d1 \<Longrightarrow> cmod (f y - f x) < e/2"
+      unfolding continuous_at Lim_at dist_norm  using e
+      by (drule_tac x="e/2" in spec) force
+    obtain d2 where d2: "d2>0" "ball x d2 \<subseteq> s" using  \<open>open s\<close> x
+      by (auto simp: open_contains_ball)
+    have dpos: "min d1 d2 > 0" using d1 d2 by simp
+    { fix y
+      assume yx: "y \<noteq> x" and close: "cmod (y - x) < min d1 d2"
+      have y: "y \<in> s"
+        using d2 close  by (force simp: dist_norm norm_minus_commute)
+      have fxy: "f contour_integrable_on linepath x y"
+        apply (rule contour_integrable_continuous_linepath)
+        apply (rule continuous_on_subset [OF contf])
+        using close d2
+        apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1))
+        done
+      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
+        by (auto simp: contour_integrable_on_def)
+      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
+        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
+      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
+        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
+        using e apply simp
+        apply (rule d1_less [THEN less_imp_le])
+        using close segment_bound
+        apply force
+        done
+      also have "... < e * cmod (y - x)"
+        by (simp add: e yx)
+      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
+        using i yx  by (simp add: contour_integral_unique divide_less_eq)
+    }
+    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"
+      using dpos by blast
+  }
+  then have *: "(\<lambda>y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \<midarrow>x\<rightarrow> 0"
+    by (simp add: Lim_at dist_norm inverse_eq_divide)
+  show ?thesis
+    apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right)
+    apply (rule Lim_transform [OF * Lim_eventually])
+    apply (simp add: inverse_eq_divide [symmetric] eventually_at)
+    using \<open>open s\<close> x
+    apply (force simp: dist_norm open_contains_ball)
+    done
+qed
+
+(** Existence of a primitive.*)
+
+lemma holomorphic_starlike_primitive:
+  fixes f :: "complex \<Rightarrow> complex"
+  assumes contf: "continuous_on s f"
+      and s: "starlike s" and os: "open s"
+      and k: "finite k"
+      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
+    shows "\<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x)"
+proof -
+  obtain a where a: "a\<in>s" and a_cs: "\<And>x. x\<in>s \<Longrightarrow> closed_segment a x \<subseteq> s"
+    using s by (auto simp: starlike_def)
+  { fix x b c
+    assume "x \<in> s" "closed_segment b c \<subseteq> s"
+    then have abcs: "convex hull {a, b, c} \<subseteq> s"
+      by (simp add: a a_cs starlike_convex_subset)
+    then have *: "continuous_on (convex hull {a, b, c}) f"
+      by (simp add: continuous_on_subset [OF contf])
+    have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
+      apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
+      using abcs apply (simp add: continuous_on_subset [OF contf])
+      using * abcs interior_subset apply (auto intro: fcd)
+      done
+  } note 0 = this
+  show ?thesis
+    apply (intro exI ballI)
+    apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption)
+    apply (metis a_cs)
+    apply (metis has_chain_integral_chain_integral3 0)
+    done
+qed
+
+lemma Cauchy_theorem_starlike:
+ "\<lbrakk>open s; starlike s; finite k; continuous_on s f;
+   \<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x;
+   valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
+   \<Longrightarrow> (f has_contour_integral 0)  g"
+  by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open)
+
+lemma Cauchy_theorem_starlike_simple:
+  "\<lbrakk>open s; starlike s; f holomorphic_on s; valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
+   \<Longrightarrow> (f has_contour_integral 0) g"
+apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI])
+apply (simp_all add: holomorphic_on_imp_continuous_on)
+apply (metis at_within_open holomorphic_on_def)
+done
+
+
+subsection\<open>Cauchy's theorem for a convex set\<close>
+
+text\<open>For a convex set we can avoid assuming openness and boundary analyticity\<close>
+
+lemma triangle_contour_integrals_convex_primitive:
+  assumes contf: "continuous_on s f"
+      and s: "a \<in> s" "convex s"
+      and x: "x \<in> s"
+      and zer: "\<And>b c. \<lbrakk>b \<in> s; c \<in> s\<rbrakk>
+                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
+                       contour_integral (linepath c a) f = 0"
+    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)"
+proof -
+  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
+  { fix y
+    assume y: "y \<in> s"
+    have cont_ayf: "continuous_on (closed_segment a y) f"
+      using s y  by (meson contf continuous_on_subset convex_contains_segment)
+    have xys: "closed_segment x y \<subseteq> s"  (*?*)
+      using convex_contains_segment s x y by auto
+    have "?pathint a y - ?pathint a x = ?pathint x y"
+      using zer [OF x y]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
+  } note [simp] = this
+  { fix e::real
+    assume e: "0 < e"
+    have cont_atx: "continuous (at x within s) f"
+      using x s contf  by (simp add: continuous_on_eq_continuous_within)
+    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"
+      unfolding continuous_within Lim_within dist_norm using e
+      by (drule_tac x="e/2" in spec) force
+    { fix y
+      assume yx: "y \<noteq> x" and close: "cmod (y - x) < d1" and y: "y \<in> s"
+      have fxy: "f contour_integrable_on linepath x y"
+        using convex_contains_segment s x y
+        by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
+      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
+        by (auto simp: contour_integrable_on_def)
+      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
+        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
+      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
+        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
+        using e apply simp
+        apply (rule d1_less [THEN less_imp_le])
+        using convex_contains_segment s(2) x y apply blast
+        using close segment_bound(1) apply fastforce
+        done
+      also have "... < e * cmod (y - x)"
+        by (simp add: e yx)
+      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
+        using i yx  by (simp add: contour_integral_unique divide_less_eq)
+    }
+    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"
+      using d1 by blast
+  }
+  then have *: "((\<lambda>y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \<longlongrightarrow> 0) (at x within s)"
+    by (simp add: Lim_within dist_norm inverse_eq_divide)
+  show ?thesis
+    apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right)
+    apply (rule Lim_transform [OF * Lim_eventually])
+    using linordered_field_no_ub
+    apply (force simp: inverse_eq_divide [symmetric] eventually_at)
+    done
+qed
+
+lemma contour_integral_convex_primitive:
+  "\<lbrakk>convex s; continuous_on s f;
+    \<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>
+         \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
+  apply (cases "s={}")
+  apply (simp_all add: ex_in_conv [symmetric])
+  apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3)
+  done
+
+lemma holomorphic_convex_primitive:
+  fixes f :: "complex \<Rightarrow> complex"
+  shows
+  "\<lbrakk>convex s; finite k; continuous_on s f;
+    \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x\<rbrakk>
+   \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
+apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite])
+prefer 3
+apply (erule continuous_on_subset)
+apply (simp add: subset_hull continuous_on_subset, assumption+)
+by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull)
+
+lemma Cauchy_theorem_convex:
+    "\<lbrakk>continuous_on s f; convex s; finite k;
+      \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x;
+     valid_path g; path_image g \<subseteq> s;
+     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
+  by (metis holomorphic_convex_primitive Cauchy_theorem_primitive)
+
+lemma Cauchy_theorem_convex_simple:
+    "\<lbrakk>f holomorphic_on s; convex s;
+     valid_path g; path_image g \<subseteq> s;
+     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
+  apply (rule Cauchy_theorem_convex)
+  apply (simp_all add: holomorphic_on_imp_continuous_on)
+  apply (rule finite.emptyI)
+  using at_within_interior holomorphic_on_def interior_subset by fastforce
+
+
+text\<open>In particular for a disc\<close>
+lemma Cauchy_theorem_disc:
+    "\<lbrakk>finite k; continuous_on (cball a e) f;
+      \<And>x. x \<in> ball a e - k \<Longrightarrow> f field_differentiable at x;
+     valid_path g; path_image g \<subseteq> cball a e;
+     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
+  apply (rule Cauchy_theorem_convex)
+  apply (auto simp: convex_cball interior_cball)
+  done
+
+lemma Cauchy_theorem_disc_simple:
+    "\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e;
+     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
+by (simp add: Cauchy_theorem_convex_simple)
+
+
+subsection\<open>Generalize integrability to local primitives\<close>
+
+lemma contour_integral_local_primitive_lemma:
+  fixes f :: "complex\<Rightarrow>complex"
+  shows
+    "\<lbrakk>g piecewise_differentiable_on {a..b};
+      \<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s);
+      \<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s\<rbrakk>
+     \<Longrightarrow> (\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b}))
+            integrable_on {a..b}"
+  apply (cases "cbox a b = {}", force)
+  apply (simp add: integrable_on_def)
+  apply (rule exI)
+  apply (rule contour_integral_primitive_lemma, assumption+)
+  using atLeastAtMost_iff by blast
+
+lemma contour_integral_local_primitive_any:
+  fixes f :: "complex \<Rightarrow> complex"
+  assumes gpd: "g piecewise_differentiable_on {a..b}"
+      and dh: "\<And>x. x \<in> s
+               \<Longrightarrow> \<exists>d h. 0 < d \<and>
+                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
+      and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
+  shows "(\<lambda>x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}"
+proof -
+  { fix x
+    assume x: "a \<le> x" "x \<le> b"
+    obtain d h where d: "0 < d"
+               and h: "(\<And>y. norm(y - g x) < d \<Longrightarrow> (h has_field_derivative f y) (at y within s))"
+      using x gs dh by (metis atLeastAtMost_iff)
+    have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast
+    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"
+      using x d
+      apply (auto simp: dist_norm continuous_on_iff)
+      apply (drule_tac x=x in bspec)
+      using x apply simp
+      apply (drule_tac x=d in spec, auto)
+      done
+    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>
+                          (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}"
+      apply (rule_tac x=e in exI)
+      using e
+      apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify)
+      apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma)
+        apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset)
+       apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force)
+      done
+  } then
+  show ?thesis
+    by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified])
+qed
+
+lemma contour_integral_local_primitive:
+  fixes f :: "complex \<Rightarrow> complex"
+  assumes g: "valid_path g" "path_image g \<subseteq> s"
+      and dh: "\<And>x. x \<in> s
+               \<Longrightarrow> \<exists>d h. 0 < d \<and>
+                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
+  shows "f contour_integrable_on g"
+  using g
+  apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def
+            has_integral_localized_vector_derivative integrable_on_def [symmetric])
+  using contour_integral_local_primitive_any [OF _ dh]
+  by (meson image_subset_iff piecewise_C1_imp_differentiable)
+
+
+text\<open>In particular if a function is holomorphic\<close>
+
+lemma contour_integrable_holomorphic:
+  assumes contf: "continuous_on s f"
+      and os: "open s"
+      and k: "finite k"
+      and g: "valid_path g" "path_image g \<subseteq> s"
+      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
+    shows "f contour_integrable_on g"
+proof -
+  { fix z
+    assume z: "z \<in> s"
+    obtain d where d: "d>0" "ball z d \<subseteq> s" using  \<open>open s\<close> z
+      by (auto simp: open_contains_ball)
+    then have contfb: "continuous_on (ball z d) f"
+      using contf continuous_on_subset by blast
+    obtain h where "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within ball z d)"
+      using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d
+            interior_subset by force
+    then have "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within s)"
+      by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE)
+    then have "\<exists>h. (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
+      by (force simp: dist_norm norm_minus_commute)
+    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))"
+      using d by blast
+  }
+  then show ?thesis
+    by (rule contour_integral_local_primitive [OF g])
+qed
+
+lemma contour_integrable_holomorphic_simple:
+  assumes fh: "f holomorphic_on s"
+      and os: "open s"
+      and g: "valid_path g" "path_image g \<subseteq> s"
+    shows "f contour_integrable_on g"
+  apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g])
+  apply (simp add: fh holomorphic_on_imp_continuous_on)
+  using fh  by (simp add: field_differentiable_def holomorphic_on_open os)
+
+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 contour_integrable_inversediff:
+    "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) contour_integrable_on g"
+apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"])
+apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros)
+done
+
+text\<open>Key fact that path integral is the same for a "nearby" path. This is the
+ main lemma for the homotopy form of Cauchy's theorem and is also useful
+ if we want "without loss of generality" to assume some nice properties of a
+ path (e.g. smoothness). It can also be used to define the integrals of
+ analytic functions over arbitrary continuous paths. This is just done for
+ winding numbers now.
+\<close>
+
+text\<open>A technical definition to avoid duplication of similar proofs,
+     for paths joined at the ends versus looping paths\<close>
+definition linked_paths :: "bool \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
+  where "linked_paths atends g h ==
+        (if atends then pathstart h = pathstart g \<and> pathfinish h = pathfinish g
+                   else pathfinish g = pathstart g \<and> pathfinish h = pathstart h)"
+
+text\<open>This formulation covers two cases: @{term g} and @{term h} share their
+      start and end points; @{term g} and @{term h} both loop upon themselves.\<close>
+lemma contour_integral_nearby:
+  assumes os: "open s" and p: "path p" "path_image p \<subseteq> s"
+    shows
+       "\<exists>d. 0 < d \<and>
+            (\<forall>g h. valid_path g \<and> valid_path h \<and>
+                  (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
+                  linked_paths atends g h
+                  \<longrightarrow> 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))"
+proof -
+  have "\<forall>z. \<exists>e. z \<in> path_image p \<longrightarrow> 0 < e \<and> ball z e \<subseteq> s"
+    using open_contains_ball os p(2) by blast
+  then obtain ee where ee: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s"
+    by metis
+  define cover where "cover = (\<lambda>z. ball z (ee z/3)) ` (path_image p)"
+  have "compact (path_image p)"
+    by (metis p(1) compact_path_image)
+  moreover have "path_image p \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))"
+    using ee by auto
+  ultimately have "\<exists>D \<subseteq> cover. finite D \<and> path_image p \<subseteq> \<Union>D"
+    by (simp add: compact_eq_heine_borel cover_def)
+  then obtain D where D: "D \<subseteq> cover" "finite D" "path_image p \<subseteq> \<Union>D"
+    by blast
+  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"
+    apply (simp add: cover_def path_image_def image_comp)
+    apply (blast dest!: finite_subset_image [OF \<open>finite D\<close>])
+    done
+  then have kne: "k \<noteq> {}"
+    using D by auto
+  have pi: "\<And>i. i \<in> k \<Longrightarrow> p i \<in> path_image p"
+    using k  by (auto simp: path_image_def)
+  then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))"
+    by (metis ee)
+  define e where "e = Min((ee o p) ` k)"
+  have fin_eep: "finite ((ee o p) ` k)"
+    using k  by blast
+  have enz: "0 < e"
+    using ee k  by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi)
+  have "uniformly_continuous_on {0..1} p"
+    using p  by (simp add: path_def compact_uniformly_continuous)
+  then obtain d::real where d: "d>0"
+          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"
+    unfolding uniformly_continuous_on_def dist_norm real_norm_def
+    by (metis divide_pos_pos enz zero_less_numeral)
+  then obtain N::nat where N: "N>0" "inverse N < d"
+    using real_arch_inverse [of d]   by auto
+  { fix g h
+    assume g: "valid_path g" and gp: "\<forall>t\<in>{0..1}. cmod (g t - p t) < e / 3"
+       and h: "valid_path h" and hp: "\<forall>t\<in>{0..1}. cmod (h t - p t) < e / 3"
+       and joins: "linked_paths atends g h"
+    { fix t::real
+      assume t: "0 \<le> t" "t \<le> 1"
+      then obtain u where u: "u \<in> k" and ptu: "p t \<in> ball(p u) (ee(p u) / 3)"
+        using \<open>path_image p \<subseteq> \<Union>D\<close> D_eq by (force simp: path_image_def)
+      then have ele: "e \<le> ee (p u)" using fin_eep
+        by (simp add: e_def)
+      have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3"
+        using gp hp t by auto
+      with ele have "cmod (g t - p t) < ee (p u) / 3"
+                    "cmod (h t - p t) < ee (p u) / 3"
+        by linarith+
+      then have "g t \<in> ball(p u) (ee(p u))"  "h t \<in> ball(p u) (ee(p u))"
+        using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"]
+              norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u
+        by (force simp: dist_norm ball_def norm_minus_commute)+
+      then have "g t \<in> s" "h t \<in> s" using ee u k
+        by (auto simp: path_image_def ball_def)
+    }
+    then have ghs: "path_image g \<subseteq> s" "path_image h \<subseteq> s"
+      by (auto simp: path_image_def)
+    moreover
+    { fix f
+      assume fhols: "f holomorphic_on s"
+      then have fpa: "f contour_integrable_on g"  "f contour_integrable_on h"
+        using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple
+        by blast+
+      have contf: "continuous_on s f"
+        by (simp add: fhols holomorphic_on_imp_continuous_on)
+      { fix z
+        assume z: "z \<in> path_image p"
+        have "f holomorphic_on ball z (ee z)"
+          using fhols ee z holomorphic_on_subset by blast
+        then have "\<exists>ff. (\<forall>w \<in> ball z (ee z). (ff has_field_derivative f w) (at w))"
+          using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified]
+          by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball)
+      }
+      then obtain ff where ff:
+            "\<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)"
+        by metis
+      { fix n
+        assume n: "n \<le> N"
+        then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f =
+                   contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f"
+        proof (induct n)
+          case 0 show ?case by simp
+        next
+          case (Suc n)
+          obtain t where t: "t \<in> k" and "p (n/N) \<in> ball(p t) (ee(p t) / 3)"
+            using \<open>path_image p \<subseteq> \<Union>D\<close> [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems
+            by (force simp: path_image_def)
+          then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3"
+            by (simp add: dist_norm)
+          have e3le: "e/3 \<le> ee (p t) / 3"  using fin_eep t
+            by (simp add: e_def)
+          { fix x
+            assume x: "n/N \<le> x" "x \<le> (1 + n)/N"
+            then have nN01: "0 \<le> n/N" "(1 + n)/N \<le> 1"
+              using Suc.prems by auto
+            then have x01: "0 \<le> x" "x \<le> 1"
+              using x by linarith+
+            have "cmod (p t - p x)  < ee (p t) / 3 + e/3"
+              apply (rule norm_diff_triangle_less [OF ptu de])
+              using x N x01 Suc.prems
+              apply (auto simp: field_simps)
+              done
+            then have ptx: "cmod (p t - p x) < 2*ee (p t)/3"
+              using e3le eepi [OF t] by simp
+            have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 "
+              apply (rule norm_diff_triangle_less [OF ptx])
+              using gp x01 by (simp add: norm_minus_commute)
+            also have "... \<le> ee (p t)"
+              using e3le eepi [OF t] by simp
+            finally have gg: "cmod (p t - g x) < ee (p t)" .
+            have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 "
+              apply (rule norm_diff_triangle_less [OF ptx])
+              using hp x01 by (simp add: norm_minus_commute)
+            also have "... \<le> ee (p t)"
+              using e3le eepi [OF t] by simp
+            finally have "cmod (p t - g x) < ee (p t)"
+                         "cmod (p t - h x) < ee (p t)"
+              using gg by auto
+          } note ptgh_ee = this
+          have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \<subseteq> ball (p t) (ee (p t))"
+            using ptgh_ee [of "n/N"] Suc.prems
+            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
+          then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \<subseteq> s"
+            using \<open>N>0\<close> Suc.prems
+            apply (simp add: path_image_join field_simps closed_segment_commute)
+            apply (erule order_trans)
+            apply (simp add: ee pi t)
+            done
+          have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N)))
+                  \<subseteq> ball (p t) (ee (p t))"
+            using ptgh_ee [of "(1+n)/N"] Suc.prems
+            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
+          then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \<subseteq> s"
+            using \<open>N>0\<close> Suc.prems ee pi t
+            by (auto simp: Path_Connected.path_image_join field_simps)
+          have pi_subset_ball:
+                "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++
+                             subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N)))
+                 \<subseteq> ball (p t) (ee (p t))"
+            apply (intro subset_path_image_join pi_hgn pi_ghn')
+            using \<open>N>0\<close> Suc.prems
+            apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee)
+            done
+          have pi0: "(f has_contour_integral 0)
+                       (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++
+                        subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))"
+            apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"])
+            apply (metis ff open_ball at_within_open pi t)
+            apply (intro valid_path_join)
+            using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h)
+            done
+          have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g"
+            using Suc.prems by (simp add: contour_integrable_subpath g fpa)
+          have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))"
+            using gh_n's
+            by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
+          have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))"
+            using gh_ns
+            by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
+          have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f +
+                     contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f +
+                     contour_integral (subpath ((Suc n) / N) (n/N) h) f +
+                     contour_integral (linepath (h (n/N)) (g (n/N))) f = 0"
+            using contour_integral_unique [OF pi0] Suc.prems
+            by (simp add: g h fpa valid_path_subpath contour_integrable_subpath
+                          fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc)
+          have *: "\<And>hn he hn' gn gd gn' hgn ghn gh0 ghn'.
+                    \<lbrakk>hn - gn = ghn - gh0;
+                     gd + ghn' + he + hgn = (0::complex);
+                     hn - he = hn'; gn + gd = gn'; hgn = -ghn\<rbrakk> \<Longrightarrow> hn' - gn' = ghn' - gh0"
+            by (auto simp: algebra_simps)
+          have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
+                contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f"
+            unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"]
+            using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath)
+          also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f"
+            using Suc.prems by (simp add: contour_integral_subpath_combine h fpa)
+          finally have pi0_eq:
+               "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
+                contour_integral (subpath 0 ((Suc n) / N) h) f" .
+          show ?case
+            apply (rule * [OF Suc.hyps eq0 pi0_eq])
+            using Suc.prems
+            apply (simp_all add: g h fpa contour_integral_subpath_combine
+                     contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath
+                     continuous_on_subset [OF contf gh_ns])
+            done
+      qed
+      } note ind = this
+      have "contour_integral h f = contour_integral g f"
+        using ind [OF order_refl] N joins
+        by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm)
+    }
+    ultimately
+    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)"
+      by metis
+  } note * = this
+  show ?thesis
+    apply (rule_tac x="e/3" in exI)
+    apply (rule conjI)
+    using enz apply simp
+    apply (clarsimp simp only: ball_conj_distrib)
+    apply (rule *; assumption)
+    done
+qed
+
+
+lemma
+  assumes "open s" "path p" "path_image p \<subseteq> s"
+    shows contour_integral_nearby_ends:
+      "\<exists>d. 0 < d \<and>
+              (\<forall>g h. valid_path g \<and> valid_path h \<and>
+                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
+                    pathstart h = pathstart g \<and> pathfinish h = pathfinish g
+                    \<longrightarrow> 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))"
+    and contour_integral_nearby_loops:
+      "\<exists>d. 0 < d \<and>
+              (\<forall>g h. valid_path g \<and> valid_path h \<and>
+                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
+                    pathfinish g = pathstart g \<and> pathfinish h = pathstart h
+                    \<longrightarrow> 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))"
+  using contour_integral_nearby [OF assms, where atends=True]
+  using contour_integral_nearby [OF assms, where atends=False]
+  unfolding linked_paths_def by simp_all
+
+corollary differentiable_polynomial_function:
+  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "polynomial_function p \<Longrightarrow> p differentiable_on s"
+by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def)
+
+lemma C1_differentiable_polynomial_function:
+  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "polynomial_function p \<Longrightarrow> p C1_differentiable_on s"
+  by (metis continuous_on_polymonial_function C1_differentiable_on_def  has_vector_derivative_polynomial_function)
+
+lemma valid_path_polynomial_function:
+  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "polynomial_function p \<Longrightarrow> valid_path p"
+by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
+
+lemma valid_path_subpath_trivial [simp]:
+    fixes g :: "real \<Rightarrow> 'a::euclidean_space"
+    shows "z \<noteq> g x \<Longrightarrow> valid_path (subpath x x g)"
+  by (simp add: subpath_def valid_path_polynomial_function)
+
+lemma contour_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(contour_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> contour_integral g f = contour_integral p f"
+  using contour_integral_nearby_ends [OF s \<open>path g\<close> 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 contour_integrable_on p \<and> valid_path p"
+    using p s
+    by (blast intro: valid_path_polynomial_function contour_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 (contour_integral g f) \<le> L * B"
+    apply (simp add: pi [OF f])
+    apply (simp add: contour_integral_integral)
+    apply (rule order_trans [OF integral_norm_bound_integral])
+    apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult)
+    done
+} then
+show ?thesis
+  by (force simp: L contour_integral_integral)
+qed
+
+subsection\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
+
+text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
+
+lemma continuous_disconnected_range_constant:
+  assumes s: "connected s"
+      and conf: "continuous_on s f"
+      and fim: "f ` s \<subseteq> t"
+      and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
+    shows "\<exists>a. \<forall>x \<in> s. f x = a"
+proof (cases "s = {}")
+  case True then show ?thesis by force
+next
+  case False
+  { fix x assume "x \<in> s"
+    then have "f ` s \<subseteq> {f x}"
+    by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct)
+  }
+  with False show ?thesis
+    by blast
+qed
+
+lemma discrete_subset_disconnected:
+  fixes s :: "'a::topological_space set"
+  fixes t :: "'b::real_normed_vector set"
+  assumes conf: "continuous_on s f"
+      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)"
+   shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}"
+proof -
+  { fix x assume x: "x \<in> s"
+    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)"
+      using conf no [OF x] by auto
+    then have e2: "0 \<le> e / 2"
+      by simp
+    have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y
+      apply (rule ccontr)
+      using connected_closed [of "connected_component_set (f ` s) (f x)"] \<open>e>0\<close>
+      apply (simp add: del: ex_simps)
+      apply (drule spec [where x="cball (f x) (e / 2)"])
+      apply (drule spec [where x="- ball(f x) e"])
+      apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in)
+        apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less)
+       using centre_in_cball connected_component_refl_eq e2 x apply blast
+      using ccs
+      apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>])
+      done
+    moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
+      by (auto simp: connected_component_in)
+    ultimately have "connected_component_set (f ` s) (f x) = {f x}"
+      by (auto simp: x)
+  }
+  with assms show ?thesis
+    by blast
+qed
+
+lemma finite_implies_discrete:
+  fixes s :: "'a::topological_space set"
+  assumes "finite (f ` s)"
+  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))"
+proof -
+  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
+  proof (cases "f ` s - {f x} = {}")
+    case True
+    with zero_less_numeral show ?thesis
+      by (fastforce simp add: Set.image_subset_iff cong: conj_cong)
+  next
+    case False
+    then obtain z where z: "z \<in> s" "f z \<noteq> f x"
+      by blast
+    have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}"
+      using assms by simp
+    then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}"
+      apply (rule finite_imp_less_Inf)
+      using z apply force+
+      done
+    show ?thesis
+      by (force intro!: * cInf_le_finite [OF finn])
+  qed
+  with assms show ?thesis
+    by blast
+qed
+
+text\<open>This proof requires the existence of two separate values of the range type.\<close>
+lemma finite_range_constant_imp_connected:
+  assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
+              \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a"
+    shows "connected s"
+proof -
+  { fix t u
+    assume clt: "closedin (subtopology euclidean s) t"
+       and clu: "closedin (subtopology euclidean s) u"
+       and tue: "t \<inter> u = {}" and tus: "t \<union> u = s"
+    have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)"
+      apply (subst tus [symmetric])
+      apply (rule continuous_on_cases_local)
+      using clt clu tue
+      apply (auto simp: tus continuous_on_const)
+      done
+    have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)"
+      by (rule finite_subset [of _ "{0,1}"]) auto
+    have "t = {} \<or> u = {}"
+      using assms [OF conif fi] tus [symmetric]
+      by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue)
+  }
+  then show ?thesis
+    by (simp add: connected_closedin_eq)
+qed
+
+lemma continuous_disconnected_range_constant_eq:
+      "(connected s \<longleftrightarrow>
+           (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
+            \<forall>t. continuous_on s f \<and> f ` s \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
+            \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis1)
+  and continuous_discrete_range_constant_eq:
+      "(connected s \<longleftrightarrow>
+         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
+          continuous_on s f \<and>
+          (\<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)))
+          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis2)
+  and continuous_finite_range_constant_eq:
+      "(connected s \<longleftrightarrow>
+         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
+          continuous_on s f \<and> finite (f ` s)
+          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis3)
+proof -
+  have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk>
+    \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)"
+    by blast
+  have "?thesis1 \<and> ?thesis2 \<and> ?thesis3"
+    apply (rule *)
+    using continuous_disconnected_range_constant apply metis
+    apply clarify
+    apply (frule discrete_subset_disconnected; blast)
+    apply (blast dest: finite_implies_discrete)
+    apply (blast intro!: finite_range_constant_imp_connected)
+    done
+  then show ?thesis1 ?thesis2 ?thesis3
+    by blast+
+qed
+
+lemma continuous_discrete_range_constant:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
+  assumes s: "connected s"
+      and "continuous_on s f"
+      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)"
+    shows "\<exists>a. \<forall>x \<in> s. f x = a"
+  using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms
+  by blast
+
+lemma continuous_finite_range_constant:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
+  assumes "connected s"
+      and "continuous_on s f"
+      and "finite (f ` s)"
+    shows "\<exists>a. \<forall>x \<in> s. f x = a"
+  using assms continuous_finite_range_constant_eq
+  by blast
+
+
+text\<open>We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\<close>
+
+subsection\<open>Winding Numbers\<close>
+
+definition winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
+  "winding_number \<gamma> z \<equiv>
+    @n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
+                    pathstart p = pathstart \<gamma> \<and>
+                    pathfinish p = pathfinish \<gamma> \<and>
+                    (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
+                    contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+
+lemma winding_number:
+  assumes "path \<gamma>" "z \<notin> path_image \<gamma>" "0 < e"
+    shows "\<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
+               pathstart p = pathstart \<gamma> \<and>
+               pathfinish p = pathfinish \<gamma> \<and>
+               (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
+               contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * winding_number \<gamma> z"
+proof -
+  have "path_image \<gamma> \<subseteq> UNIV - {z}"
+    using assms by blast
+  then obtain d
+    where d: "d>0"
+      and pi_eq: "\<And>h1 h2. valid_path h1 \<and> valid_path h2 \<and>
+                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d) \<and>
+                    pathstart h2 = pathstart h1 \<and> pathfinish h2 = pathfinish h1 \<longrightarrow>
+                      path_image h1 \<subseteq> UNIV - {z} \<and> path_image h2 \<subseteq> UNIV - {z} \<and>
+                      (\<forall>f. f holomorphic_on UNIV - {z} \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
+    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete)
+  then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
+                          (\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
+    using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto
+  define nn where "nn = 1/(2* pi*\<i>) * contour_integral h (\<lambda>w. 1/(w - z))"
+  have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
+                        pathstart p = pathstart \<gamma> \<and>  pathfinish p = pathfinish \<gamma> \<and>
+                        (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
+                        contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+                    (is "\<exists>n. \<forall>e > 0. ?PP e n")
+    proof (rule_tac x=nn in exI, clarify)
+      fix e::real
+      assume e: "e>0"
+      obtain p where p: "polynomial_function p \<and>
+            pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))"
+        using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "min e (d/2)"] d \<open>0<e\<close> by auto
+      have "(\<lambda>w. 1 / (w - z)) holomorphic_on UNIV - {z}"
+        by (auto simp: intro!: holomorphic_intros)
+      then show "?PP e nn"
+        apply (rule_tac x=p in exI)
+        using pi_eq [of h p] h p d
+        apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def)
+        done
+    qed
+  then show ?thesis
+    unfolding winding_number_def
+    apply (rule someI2_ex)
+    apply (blast intro: \<open>0<e\<close>)
+    done
+qed
+
+lemma winding_number_unique:
+  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
+      and pi:
+        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
+                          pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
+                          (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
+                          contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+   shows "winding_number \<gamma> z = n"
+proof -
+  have "path_image \<gamma> \<subseteq> UNIV - {z}"
+    using assms by blast
+  then obtain e
+    where e: "e>0"
+      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
+                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
+                    pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
+                    contour_integral h2 f = contour_integral h1 f"
+    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
+  obtain p where p:
+     "valid_path p \<and> z \<notin> path_image p \<and>
+      pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
+      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
+      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+    using pi [OF e] by blast
+  obtain q where q:
+     "valid_path q \<and> z \<notin> path_image q \<and>
+      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
+      (\<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"
+    using winding_number [OF \<gamma> e] by blast
+  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
+    using p by auto
+  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
+    apply (rule pi_eq)
+    using p q
+    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
+  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
+    using q by auto
+  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
+  then show ?thesis
+    by simp
+qed
+
+lemma winding_number_unique_loop:
+  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
+      and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+      and pi:
+        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
+                           pathfinish p = pathstart p \<and>
+                           (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
+                           contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+   shows "winding_number \<gamma> z = n"
+proof -
+  have "path_image \<gamma> \<subseteq> UNIV - {z}"
+    using assms by blast
+  then obtain e
+    where e: "e>0"
+      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
+                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
+                    pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
+                    contour_integral h2 f = contour_integral h1 f"
+    using contour_integral_nearby_loops [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
+  obtain p where p:
+     "valid_path p \<and> z \<notin> path_image p \<and>
+      pathfinish p = pathstart p \<and>
+      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
+      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
+    using pi [OF e] by blast
+  obtain q where q:
+     "valid_path q \<and> z \<notin> path_image q \<and>
+      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
+      (\<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"
+    using winding_number [OF \<gamma> e] by blast
+  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
+    using p by auto
+  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
+    apply (rule pi_eq)
+    using p q loop
+    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
+  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
+    using q by auto
+  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
+  then show ?thesis
+    by simp
+qed
+
+lemma winding_number_valid_path:
+  assumes "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
+    shows "winding_number \<gamma> z = 1/(2*pi*\<i>) * contour_integral \<gamma> (\<lambda>w. 1/(w - z))"
+  using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique)
+
+lemma has_contour_integral_winding_number:
+  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
+    shows "((\<lambda>w. 1/(w - z)) has_contour_integral (2*pi*\<i>*winding_number \<gamma> z)) \<gamma>"
+by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms)
+
+lemma winding_number_trivial [simp]: "z \<noteq> a \<Longrightarrow> winding_number(linepath a a) z = 0"
+  by (simp add: winding_number_valid_path)
+
+lemma winding_number_subpath_trivial [simp]: "z \<noteq> g x \<Longrightarrow> winding_number (subpath x x g) z = 0"
+  by (simp add: path_image_subpath winding_number_valid_path)
+
+lemma winding_number_join:
+  assumes g1: "path g1" "z \<notin> path_image g1"
+      and g2: "path g2" "z \<notin> path_image g2"
+      and "pathfinish g1 = pathstart g2"
+    shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z"
+  apply (rule winding_number_unique)
+  using assms apply (simp_all add: not_in_path_image_join)
+  apply (frule winding_number [OF g2])
+  apply (frule winding_number [OF g1], clarify)
+  apply (rename_tac p2 p1)
+  apply (rule_tac x="p1+++p2" in exI)
+  apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps)
+  apply (auto simp: joinpaths_def)
+  done
+
+lemma winding_number_reversepath:
+  assumes "path \<gamma>" "z \<notin> path_image \<gamma>"
+    shows "winding_number(reversepath \<gamma>) z = - (winding_number \<gamma> z)"
+  apply (rule winding_number_unique)
+  using assms
+  apply simp_all
+  apply (frule winding_number [OF assms], clarify)
+  apply (rule_tac x="reversepath p" in exI)
+  apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse)
+  apply (auto simp: reversepath_def)
+  done
+
+lemma winding_number_shiftpath:
+  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
+      and "pathfinish \<gamma> = pathstart \<gamma>" "a \<in> {0..1}"
+    shows "winding_number(shiftpath a \<gamma>) z = winding_number \<gamma> z"
+  apply (rule winding_number_unique_loop)
+  using assms
+  apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath)
+  apply (frule winding_number [OF \<gamma>], clarify)
+  apply (rule_tac x="shiftpath a p" in exI)
+  apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath)
+  apply (auto simp: shiftpath_def)
+  done
+
+lemma winding_number_split_linepath:
+  assumes "c \<in> closed_segment a b" "z \<notin> closed_segment a b"
+    shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z"
+proof -
+  have "z \<notin> closed_segment a c" "z \<notin> closed_segment c b"
+    using assms  apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE)
+    using assms  by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE)
+  then show ?thesis
+    using assms
+    by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps)
+qed
+
+lemma winding_number_cong:
+   "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p t = q t) \<Longrightarrow> winding_number p z = winding_number q z"
+  by (simp add: winding_number_def pathstart_def pathfinish_def)
+
+lemma winding_number_offset: "winding_number p z = winding_number (\<lambda>w. p w - z) 0"
+  apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def)
+  apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe)
+  apply (rename_tac g)
+  apply (rule_tac x="\<lambda>t. g t - z" in exI)
+  apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise)
+  apply (rename_tac g)
+  apply (rule_tac x="\<lambda>t. g t + z" in exI)
+  apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise)
+  apply (force simp: algebra_simps)
+  done
+
+(* A combined theorem deducing several things piecewise.*)
+lemma winding_number_join_pos_combined:
+     "\<lbrakk>valid_path \<gamma>1; z \<notin> path_image \<gamma>1; 0 < Re(winding_number \<gamma>1 z);
+       valid_path \<gamma>2; z \<notin> path_image \<gamma>2; 0 < Re(winding_number \<gamma>2 z); pathfinish \<gamma>1 = pathstart \<gamma>2\<rbrakk>
+      \<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)"
+  by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path)
+
+
+(* Useful sufficient conditions for the winding number to be positive etc.*)
+
+lemma Re_winding_number:
+    "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk>
+     \<Longrightarrow> Re(winding_number \<gamma> z) = Im(contour_integral \<gamma> (\<lambda>w. 1/(w - z))) / (2*pi)"
+by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square)
+
+lemma winding_number_pos_le:
+  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
+      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> 0 \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
+    shows "0 \<le> Re(winding_number \<gamma> z)"
+proof -
+  have *: "0 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" if x: "0 < x" "x < 1" for x
+    using ge by (simp add: Complex.Im_divide algebra_simps x)
+  show ?thesis
+    apply (simp add: Re_winding_number [OF \<gamma>] field_simps)
+    apply (rule has_integral_component_nonneg
+             [of \<i> "\<lambda>x. if x \<in> {0<..<1}
+                         then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else 0", simplified])
+      prefer 3 apply (force simp: *)
+     apply (simp add: Basis_complex_def)
+    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)"])
+    apply simp
+    apply (simp only: box_real)
+    apply (subst has_contour_integral [symmetric])
+    using \<gamma>
+    apply (simp add: contour_integrable_inversediff has_contour_integral_integral)
+    done
+qed
+
+lemma winding_number_pos_lt_lemma:
+  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
+      and e: "0 < e"
+      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
+    shows "0 < Re(winding_number \<gamma> z)"
+proof -
+  have "e \<le> Im (contour_integral \<gamma> (\<lambda>w. 1 / (w - z)))"
+    apply (rule has_integral_component_le
+             [of \<i> "\<lambda>x. \<i>*e" "\<i>*e" "{0..1}"
+                    "\<lambda>x. if x \<in> {0<..<1} then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else \<i>*e"
+                    "contour_integral \<gamma> (\<lambda>w. 1/(w - z))", simplified])
+    using e
+    apply (simp_all add: Basis_complex_def)
+    using has_integral_const_real [of _ 0 1] apply force
+    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)", simplified box_real])
+    apply simp
+    apply (subst has_contour_integral [symmetric])
+    using \<gamma>
+    apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge)
+    done
+  with e show ?thesis
+    by (simp add: Re_winding_number [OF \<gamma>] field_simps)
+qed
+
+lemma winding_number_pos_lt:
+  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
+      and e: "0 < e"
+      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
+    shows "0 < Re (winding_number \<gamma> z)"
+proof -
+  have bm: "bounded ((\<lambda>w. w - z) ` (path_image \<gamma>))"
+    using bounded_translation [of _ "-z"] \<gamma> by (simp add: bounded_valid_path_image)
+  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"
+    using bounded_pos [THEN iffD1, OF bm] by blast
+  { fix x::real  assume x: "0 < x" "x < 1"
+    then have B2: "cmod (\<gamma> x - z)^2 \<le> B^2" using Bno [of "\<gamma> x - z"]
+      by (simp add: path_image_def power2_eq_square mult_mono')
+    with x have "\<gamma> x \<noteq> z" using \<gamma>
+      using path_image_def by fastforce
+    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) * cnj (\<gamma> x - z)) / (cmod (\<gamma> x - z))\<^sup>2"
+      using B ge [OF x] B2 e
+      apply (rule_tac y="e / (cmod (\<gamma> x - z))\<^sup>2" in order_trans)
+      apply (auto simp: divide_left_mono divide_right_mono)
+      done
+    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
+      by (simp add: Im_divide_Reals complex_div_cnj [of _ "\<gamma> x - z" for x] del: complex_cnj_diff times_complex.sel)
+  } note * = this
+  show ?thesis
+    using e B by (simp add: * winding_number_pos_lt_lemma [OF \<gamma>, of "e/B^2"])
+qed
+
+subsection\<open>The winding number is an integer\<close>
+
+text\<open>Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1,
+     Also on page 134 of Serge Lang's book with the name title, etc.\<close>
+
+lemma exp_fg:
+  fixes z::complex
+  assumes g: "(g has_vector_derivative g') (at x within s)"
+      and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)"
+      and z: "g x \<noteq> z"
+    shows "((\<lambda>x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)"
+proof -
+  have *: "(exp o (\<lambda>x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)"
+    using assms unfolding has_vector_derivative_def scaleR_conv_of_real
+    by (auto intro!: derivative_eq_intros)
+  show ?thesis
+    apply (rule has_vector_derivative_eq_rhs)
+    apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult])
+    using z
+    apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g)
+    done
+qed
+
+lemma winding_number_exp_integral:
+  fixes z::complex
+  assumes \<gamma>: "\<gamma> piecewise_C1_differentiable_on {a..b}"
+      and ab: "a \<le> b"
+      and z: "z \<notin> \<gamma> ` {a..b}"
+    shows "(\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)) integrable_on {a..b}"
+          (is "?thesis1")
+          "exp (- (integral {a..b} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))) * (\<gamma> b - z) = \<gamma> a - z"
+          (is "?thesis2")
+proof -
+  let ?D\<gamma> = "\<lambda>x. vector_derivative \<gamma> (at x)"
+  have [simp]: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<gamma> x \<noteq> z"
+    using z by force
+  have cong: "continuous_on {a..b} \<gamma>"
+    using \<gamma> by (simp add: piecewise_C1_differentiable_on_def)
+  obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)"
+    using \<gamma> by (force simp: piecewise_C1_differentiable_on_def)
+  have o: "open ({a<..<b} - k)"
+    using \<open>finite k\<close> by (simp add: finite_imp_closed open_Diff)
+  moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
+    by force
+  ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
+    by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open)
+  { fix w
+    assume "w \<noteq> z"
+    have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))"
+      by (auto simp: dist_norm intro!: continuous_intros)
+    moreover have "\<And>x. cmod (w - x) < cmod (w - z) \<Longrightarrow> \<exists>f'. ((\<lambda>w. 1 / (w - z)) has_field_derivative f') (at x)"
+      by (auto simp: intro!: derivative_eq_intros)
+    ultimately have "\<exists>h. \<forall>y. norm(y - w) < norm(w - z) \<longrightarrow> (h has_field_derivative 1/(y - z)) (at y)"
+      using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\<lambda>w. 1/(w - z)"]
+      by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute)
+  }
+  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)"
+    by meson
+  have exy: "\<exists>y. ((\<lambda>x. inverse (\<gamma> x - z) * ?D\<gamma> x) has_integral y) {a..b}"
+    unfolding integrable_on_def [symmetric]
+    apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \<gamma>], of "-{z}"])
+    apply (rename_tac w)
+    apply (rule_tac x="norm(w - z)" in exI)
+    apply (simp_all add: inverse_eq_divide)
+    apply (metis has_field_derivative_at_within h)
+    done
+  have vg_int: "(\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)) integrable_on {a..b}"
+    unfolding box_real [symmetric] divide_inverse_commute
+    by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab)
+  with ab show ?thesis1
+    by (simp add: divide_inverse_commute integral_def integrable_on_def)
+  { fix t
+    assume t: "t \<in> {a..b}"
+    have cball: "continuous_on (ball (\<gamma> t) (dist (\<gamma> t) z)) (\<lambda>x. inverse (x - z))"
+        using z by (auto intro!: continuous_intros simp: dist_norm)
+    have icd: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> (\<lambda>w. inverse (w - z)) field_differentiable at x"
+      unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros)
+    obtain h where h: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow>
+                       (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\<gamma> t - y) < cmod (\<gamma> t - z)})"
+      using holomorphic_convex_primitive [where f = "\<lambda>w. inverse(w - z)", OF convex_ball finite.emptyI cball icd]
+      by simp (auto simp: ball_def dist_norm that)
+    { fix x D
+      assume x: "x \<notin> k" "a < x" "x < b"
+      then have "x \<in> interior ({a..b} - k)"
+        using open_subset_interior [OF o] by fastforce
+      then have con: "isCont (\<lambda>x. ?D\<gamma> x) x"
+        using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior)
+      then have con_vd: "continuous (at x within {a..b}) (\<lambda>x. ?D\<gamma> x)"
+        by (rule continuous_at_imp_continuous_within)
+      have gdx: "\<gamma> differentiable at x"
+        using x by (simp add: g_diff_at)
+      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))
+          (at x within {a..b})"
+        using x gdx t
+        apply (clarsimp simp add: differentiable_iff_scaleR)
+        apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within)
+        apply (simp_all add: has_vector_derivative_def [symmetric])
+        apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at])
+        apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+
+        done
+      } note * = this
+    have "exp (- (integral {a..t} (\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)))) * (\<gamma> t - z) =\<gamma> a - z"
+      apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \<union> k" a b])
+      using t
+      apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int]  simp add: ab)+
+      done
+   }
+  with ab show ?thesis2
+    by (simp add: divide_inverse_commute integral_def)
+qed
+
+corollary winding_number_exp_2pi:
+    "\<lbrakk>path p; z \<notin> path_image p\<rbrakk>
+     \<Longrightarrow> pathfinish p - z = exp (2 * pi * \<i> * winding_number p z) * (pathstart p - z)"
+using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def
+  by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus)
+
+
+subsection\<open>The version with complex integers and equality\<close>
+
+lemma integer_winding_number_eq:
+  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
+  shows "winding_number \<gamma> z \<in> \<int> \<longleftrightarrow> pathfinish \<gamma> = pathstart \<gamma>"
+proof -
+  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)"
+      by (simp add: field_simps) algebra
+  obtain p where p: "valid_path p" "z \<notin> path_image p"
+                    "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
+                    "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
+    using winding_number [OF assms, of 1] by auto
+  have [simp]: "(winding_number \<gamma> z \<in> \<int>) = (exp (contour_integral p (\<lambda>w. 1 / (w - z))) = 1)"
+      using p by (simp add: exp_eq_1 complex_is_Int_iff)
+  have "winding_number p z \<in> \<int> \<longleftrightarrow> pathfinish p = pathstart p"
+    using p z
+    apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def)
+    using winding_number_exp_integral(2) [of p 0 1 z]
+    apply (simp add: field_simps contour_integral_integral exp_minus)
+    apply (rule *)
+    apply (auto simp: path_image_def field_simps)
+    done
+  then show ?thesis using p
+    by (auto simp: winding_number_valid_path)
+qed
+
+theorem integer_winding_number:
+  "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> \<int>"
+by (metis integer_winding_number_eq)
+
+
+text\<open>If the winding number's magnitude is at least one, then the path must contain points in every direction.*)
+   We can thus bound the winding number of a path that doesn't intersect a given ray. \<close>
+
+lemma winding_number_pos_meets:
+  fixes z::complex
+  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and 1: "Re (winding_number \<gamma> z) \<ge> 1"
+      and w: "w \<noteq> z"
+  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
+proof -
+  have [simp]: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<gamma> x \<noteq> z"
+    using z by (auto simp: path_image_def)
+  have [simp]: "z \<notin> \<gamma> ` {0..1}"
+    using path_image_def z by auto
+  have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}"
+    using \<gamma> valid_path_def by blast
+  define r where "r = (w - z) / (\<gamma> 0 - z)"
+  have [simp]: "r \<noteq> 0"
+    using w z by (auto simp: r_def)
+  have "Arg r \<le> 2*pi"
+    by (simp add: Arg less_eq_real_def)
+  also have "... \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))"
+    using 1
+    apply (simp add: winding_number_valid_path [OF \<gamma> z] contour_integral_integral)
+    apply (simp add: Complex.Re_divide field_simps power2_eq_square)
+    done
+  finally have "Arg r \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" .
+  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"
+    apply (simp add:)
+    apply (rule IVT')
+    apply (simp_all add: Arg_ge_0)
+    apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp)
+    done
+  then obtain t where t:     "t \<in> {0..1}"
+                  and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
+    by blast
+  define i where "i = integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
+  have iArg: "Arg r = Im i"
+    using eqArg by (simp add: i_def)
+  have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}"
+    by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t)
+  have "exp (- i) * (\<gamma> t - z) = \<gamma> 0 - z"
+    unfolding i_def
+    apply (rule winding_number_exp_integral [OF gpdt])
+    using t z unfolding path_image_def
+    apply force+
+    done
+  then have *: "\<gamma> t - z = exp i * (\<gamma> 0 - z)"
+    by (simp add: exp_minus field_simps)
+  then have "(w - z) = r * (\<gamma> 0 - z)"
+    by (simp add: r_def)
+  then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \<gamma> t"
+    apply (simp add:)
+    apply (subst Complex_Transcendental.Arg_eq [of r])
+    apply (simp add: iArg)
+    using *
+    apply (simp add: exp_eq_polar field_simps)
+    done
+  with t show ?thesis
+    by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def)
+qed
+
+lemma winding_number_big_meets:
+  fixes z::complex
+  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "\<bar>Re (winding_number \<gamma> z)\<bar> \<ge> 1"
+      and w: "w \<noteq> z"
+  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
+proof -
+  { assume "Re (winding_number \<gamma> z) \<le> - 1"
+    then have "Re (winding_number (reversepath \<gamma>) z) \<ge> 1"
+      by (simp add: \<gamma> valid_path_imp_path winding_number_reversepath z)
+    moreover have "valid_path (reversepath \<gamma>)"
+      using \<gamma> valid_path_imp_reverse by auto
+    moreover have "z \<notin> path_image (reversepath \<gamma>)"
+      by (simp add: z)
+    ultimately have "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image (reversepath \<gamma>)"
+      using winding_number_pos_meets w by blast
+    then have ?thesis
+      by simp
+  }
+  then show ?thesis
+    using assms
+    by (simp add: abs_if winding_number_pos_meets split: if_split_asm)
+qed
+
+lemma winding_number_less_1:
+  fixes z::complex
+  shows
+  "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>; w \<noteq> z;
+    \<And>a::real. 0 < a \<Longrightarrow> z + a*(w - z) \<notin> path_image \<gamma>\<rbrakk>
+   \<Longrightarrow> \<bar>Re(winding_number \<gamma> z)\<bar> < 1"
+   by (auto simp: not_less dest: winding_number_big_meets)
+
+text\<open>One way of proving that WN=1 for a loop.\<close>
+lemma winding_number_eq_1:
+  fixes z::complex
+  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+      and 0: "0 < Re(winding_number \<gamma> z)" and 2: "Re(winding_number \<gamma> z) < 2"
+  shows "winding_number \<gamma> z = 1"
+proof -
+  have "winding_number \<gamma> z \<in> Ints"
+    by (simp add: \<gamma> integer_winding_number loop valid_path_imp_path z)
+  then show ?thesis
+    using 0 2 by (auto simp: Ints_def)
+qed
+
+
+subsection\<open>Continuity of winding number and invariance on connected sets.\<close>
+
+lemma continuous_at_winding_number:
+  fixes z::complex
+  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
+  shows "continuous (at z) (winding_number \<gamma>)"
+proof -
+  obtain e where "e>0" and cbg: "cball z e \<subseteq> - path_image \<gamma>"
+    using open_contains_cball [of "- path_image \<gamma>"]  z
+    by (force simp: closed_def [symmetric] closed_path_image [OF \<gamma>])
+  then have ppag: "path_image \<gamma> \<subseteq> - cball z (e/2)"
+    by (force simp: cball_def dist_norm)
+  have oc: "open (- cball z (e / 2))"
+    by (simp add: closed_def [symmetric])
+  obtain d where "d>0" and pi_eq:
+    "\<And>h1 h2. \<lbrakk>valid_path h1; valid_path h2;
+              (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d);
+              pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\<rbrakk>
+             \<Longrightarrow>
+               path_image h1 \<subseteq> - cball z (e / 2) \<and>
+               path_image h2 \<subseteq> - cball z (e / 2) \<and>
+               (\<forall>f. f holomorphic_on - cball z (e / 2) \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
+    using contour_integral_nearby_ends [OF oc \<gamma> ppag] by metis
+  obtain p where p: "valid_path p" "z \<notin> path_image p"
+                    "pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>"
+              and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2"
+              and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
+    using winding_number [OF \<gamma> z, of "min d e / 2"] \<open>d>0\<close> \<open>e>0\<close> by auto
+  { fix w
+    assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2"
+    then have wnotp: "w \<notin> path_image p"
+      using cbg \<open>d>0\<close> \<open>e>0\<close>
+      apply (simp add: path_image_def cball_def dist_norm, clarify)
+      apply (frule pg)
+      apply (drule_tac c="\<gamma> x" in subsetD)
+      apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l)
+      done
+    have wnotg: "w \<notin> path_image \<gamma>"
+      using cbg e2 \<open>e>0\<close> by (force simp: dist_norm norm_minus_commute)
+    { fix k::real
+      assume k: "k>0"
+      then obtain q where q: "valid_path q" "w \<notin> path_image q"
+                             "pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>"
+                    and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2"
+                    and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
+        using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] \<open>d>0\<close> \<open>e>0\<close> k
+        by (force simp: min_divide_distrib_right)
+      have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>u. 1 / (u - w))"
+        apply (rule pi_eq [OF \<open>valid_path q\<close> \<open>valid_path p\<close>, THEN conjunct2, THEN conjunct2, rule_format])
+        apply (frule pg)
+        apply (frule qg)
+        using p q \<open>d>0\<close> e2
+        apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
+        done
+      then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
+        by (simp add: pi qi)
+    } note pip = this
+    have "path p"
+      using p by (simp add: valid_path_imp_path)
+    then have "winding_number p w = winding_number \<gamma> w"
+      apply (rule winding_number_unique [OF _ wnotp])
+      apply (rule_tac x=p in exI)
+      apply (simp add: p wnotp min_divide_distrib_right pip)
+      done
+  } note wnwn = this
+  obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \<subseteq> - path_image p"
+    using p open_contains_cball [of "- path_image p"]
+    by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path])
+  obtain L
+    where "L>0"
+      and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe);
+                      \<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
+                      cmod (contour_integral p f) \<le> L * B"
+    using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \<open>valid_path p\<close> by force
+  { fix e::real and w::complex
+    assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)"
+    then have [simp]: "w \<notin> path_image p"
+      using cbp p(2) \<open>0 < pe\<close>
+      by (force simp: dist_norm norm_minus_commute path_image_def cball_def)
+    have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) =
+                  contour_integral p (\<lambda>x. 1/(x - w) - 1/(x - z))"
+      by (simp add: p contour_integrable_inversediff contour_integral_diff)
+    { fix x
+      assume pe: "3/4 * pe < cmod (z - x)"
+      have "cmod (w - x) < pe/4 + cmod (z - x)"
+        by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1))
+      then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp
+      have "cmod (z - x) \<le> cmod (z - w) + cmod (w - x)"
+        using norm_diff_triangle_le by blast
+      also have "... < pe/4 + cmod (w - x)"
+        using w by (simp add: norm_minus_commute)
+      finally have "pe/2 < cmod (w - x)"
+        using pe by auto
+      then have "(pe/2)^2 < cmod (w - x) ^ 2"
+        apply (rule power_strict_mono)
+        using \<open>pe>0\<close> by auto
+      then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2"
+        by (simp add: power_divide)
+      have "8 * L * cmod (w - z) < e * pe\<^sup>2"
+        using w \<open>L>0\<close> by (simp add: field_simps)
+      also have "... < e * 4 * cmod (w - x) * cmod (w - x)"
+        using pe2 \<open>e>0\<close> by (simp add: power2_eq_square)
+      also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))"
+        using wx
+        apply (rule mult_strict_left_mono)
+        using pe2 e not_less_iff_gr_or_eq by fastforce
+      finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)"
+        by simp
+      also have "... \<le> e * cmod (w - x) * cmod (z - x)"
+         using e by simp
+      finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" .
+      have "L * cmod (1 / (x - w) - 1 / (x - z)) \<le> e"
+        apply (cases "x=z \<or> x=w")
+        using pe \<open>pe>0\<close> w \<open>L>0\<close>
+        apply (force simp: norm_minus_commute)
+        using wx w(2) \<open>L>0\<close> pe pe2 Lwz
+        apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square)
+        done
+    } note L_cmod_le = this
+    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)"
+      apply (rule L)
+      using \<open>pe>0\<close> w
+      apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
+      using \<open>pe>0\<close> w \<open>L>0\<close>
+      apply (auto simp: cball_def dist_norm field_simps L_cmod_le  simp del: less_divide_eq_numeral1 le_divide_eq_numeral1)
+      done
+    have "cmod (contour_integral p (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e"
+      apply (simp add:)
+      apply (rule le_less_trans [OF *])
+      using \<open>L>0\<close> e
+      apply (force simp: field_simps)
+      done
+    then have "cmod (winding_number p w - winding_number p z) < e"
+      using pi_ge_two e
+      by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans)
+  } note cmod_wn_diff = this
+  then have "isCont (winding_number p) z"
+    apply (simp add: continuous_at_eps_delta, clarify)
+    apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI)
+    using \<open>pe>0\<close> \<open>L>0\<close>
+    apply (simp add: dist_norm cmod_wn_diff)
+    done
+  then show ?thesis
+    apply (rule continuous_transform_within [where d = "min d e / 2"])
+    apply (auto simp: \<open>d>0\<close> \<open>e>0\<close> dist_norm wnwn)
+    done
+qed
+
+corollary continuous_on_winding_number:
+    "path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)"
+  by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number)
+
+
+subsection\<open>The winding number is constant on a connected region\<close>
+
+lemma winding_number_constant:
+  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}"
+    shows "\<exists>k. \<forall>z \<in> s. winding_number \<gamma> z = k"
+proof -
+  { fix y z
+    assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z"
+    assume "y \<in> s" "z \<in> s"
+    then have "winding_number \<gamma> y \<in> \<int>"  "winding_number \<gamma> z \<in>  \<int>"
+      using integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> s\<close> by auto
+    with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> z)"
+      by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff)
+  } note * = this
+  show ?thesis
+    apply (rule continuous_discrete_range_constant [OF cs])
+    using continuous_on_winding_number [OF \<gamma>] sg
+    apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset)
+    apply (rule_tac x=1 in exI)
+    apply (auto simp: *)
+    done
+qed
+
+lemma winding_number_eq:
+     "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; w \<in> s; z \<in> s; connected s; s \<inter> path_image \<gamma> = {}\<rbrakk>
+      \<Longrightarrow> winding_number \<gamma> w = winding_number \<gamma> z"
+using winding_number_constant by fastforce
+
+lemma open_winding_number_levelsets:
+  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+    shows "open {z. z \<notin> path_image \<gamma> \<and> winding_number \<gamma> z = k}"
+proof -
+  have op: "open (- path_image \<gamma>)"
+    by (simp add: closed_path_image \<gamma> open_Compl)
+  { fix z assume z: "z \<notin> path_image \<gamma>" and k: "k = winding_number \<gamma> z"
+    obtain e where e: "e>0" "ball z e \<subseteq> - path_image \<gamma>"
+      using open_contains_ball [of "- path_image \<gamma>"] op z
+      by blast
+    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"
+      apply (rule_tac x=e in exI)
+      using e apply (simp add: dist_norm ball_def norm_minus_commute)
+      apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"])
+      done
+  } then
+  show ?thesis
+    by (auto simp: open_dist)
+qed
+
+subsection\<open>Winding number is zero "outside" a curve, in various senses\<close>
+
+lemma winding_number_zero_in_outside:
+  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)"
+    shows "winding_number \<gamma> z = 0"
+proof -
+  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
+    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
+  obtain w::complex where w: "w \<notin> ball 0 (B + 1)"
+    by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real)
+  have "- ball 0 (B + 1) \<subseteq> outside (path_image \<gamma>)"
+    apply (rule outside_subset_convex)
+    using B subset_ball by auto
+  then have wout: "w \<in> outside (path_image \<gamma>)"
+    using w by blast
+  moreover obtain k where "\<And>z. z \<in> outside (path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = k"
+    using winding_number_constant [OF \<gamma> loop, of "outside(path_image \<gamma>)"] connected_outside
+    by (metis DIM_complex bounded_path_image dual_order.refl \<gamma> outside_no_overlap)
+  ultimately have "winding_number \<gamma> z = winding_number \<gamma> w"
+    using z by blast
+  also have "... = 0"
+  proof -
+    have wnot: "w \<notin> path_image \<gamma>"  using wout by (simp add: outside_def)
+    { fix e::real assume "0<e"
+      obtain p where p: "polynomial_function p" "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
+                 and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)"
+                 and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)"
+        using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] \<open>e>0\<close> by force
+      have pip: "path_image p \<subseteq> ball 0 (B + 1)"
+        using B
+        apply (clarsimp simp add: path_image_def dist_norm ball_def)
+        apply (frule (1) pg1)
+        apply (fastforce dest: norm_add_less)
+        done
+      then have "w \<notin> path_image p"  using w by blast
+      then have "\<exists>p. valid_path p \<and> w \<notin> path_image p \<and>
+                     pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
+                     (\<forall>t\<in>{0..1}. cmod (\<gamma> t - p t) < e) \<and> contour_integral p (\<lambda>wa. 1 / (wa - w)) = 0"
+        apply (rule_tac x=p in exI)
+        apply (simp add: p valid_path_polynomial_function)
+        apply (intro conjI)
+        using pge apply (simp add: norm_minus_commute)
+        apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]])
+        apply (rule holomorphic_intros | simp add: dist_norm)+
+        using mem_ball_0 w apply blast
+        using p apply (simp_all add: valid_path_polynomial_function loop pip)
+        done
+    }
+    then show ?thesis
+      by (auto intro: winding_number_unique [OF \<gamma>] simp add: wnot)
+  qed
+  finally show ?thesis .
+qed
+
+lemma winding_number_zero_outside:
+    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
+  by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
+
+lemma winding_number_zero_at_infinity:
+  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+    shows "\<exists>B. \<forall>z. B \<le> norm z \<longrightarrow> winding_number \<gamma> z = 0"
+proof -
+  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
+    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
+  then show ?thesis
+    apply (rule_tac x="B+1" in exI, clarify)
+    apply (rule winding_number_zero_outside [OF \<gamma> convex_cball [of 0 B] loop])
+    apply (meson less_add_one mem_cball_0 not_le order_trans)
+    using ball_subset_cball by blast
+qed
+
+lemma winding_number_zero_point:
+    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; open s; path_image \<gamma> \<subseteq> s\<rbrakk>
+     \<Longrightarrow> \<exists>z. z \<in> s \<and> winding_number \<gamma> z = 0"
+  using outside_compact_in_open [of "path_image \<gamma>" s] path_image_nonempty winding_number_zero_in_outside
+  by (fastforce simp add: compact_path_image)
+
+
+text\<open>If a path winds round a set, it winds rounds its inside.\<close>
+lemma winding_number_around_inside:
+  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+      and cls: "closed s" and cos: "connected s" and s_disj: "s \<inter> path_image \<gamma> = {}"
+      and z: "z \<in> s" and wn_nz: "winding_number \<gamma> z \<noteq> 0" and w: "w \<in> s \<union> inside s"
+    shows "winding_number \<gamma> w = winding_number \<gamma> z"
+proof -
+  have ssb: "s \<subseteq> inside(path_image \<gamma>)"
+  proof
+    fix x :: complex
+    assume "x \<in> s"
+    hence "x \<notin> path_image \<gamma>"
+      by (meson disjoint_iff_not_equal s_disj)
+    thus "x \<in> inside (path_image \<gamma>)"
+      using \<open>x \<in> s\<close> 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)
+qed
+  show ?thesis
+    apply (rule winding_number_eq [OF \<gamma> loop w])
+    using z apply blast
+    apply (simp add: cls connected_with_inside cos)
+    apply (simp add: Int_Un_distrib2 s_disj, safe)
+    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)
+ qed
+
+
+text\<open>Bounding a WN by 1/2 for a path and point in opposite halfspaces.\<close>
+lemma winding_number_subpath_continuous:
+  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
+    shows "continuous_on {0..1} (\<lambda>x. winding_number(subpath 0 x \<gamma>) z)"
+proof -
+  have *: "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
+         winding_number (subpath 0 x \<gamma>) z"
+         if x: "0 \<le> x" "x \<le> 1" for x
+  proof -
+    have "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
+          1 / (2*pi*\<i>) * contour_integral (subpath 0 x \<gamma>) (\<lambda>w. 1/(w - z))"
+      using assms x
+      apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff])
+      done
+    also have "... = winding_number (subpath 0 x \<gamma>) z"
+      apply (subst winding_number_valid_path)
+      using assms x
+      apply (simp_all add: path_image_subpath valid_path_subpath)
+      by (force simp: path_image_def)
+    finally show ?thesis .
+  qed
+  show ?thesis
+    apply (rule continuous_on_eq
+                 [where f = "\<lambda>x. 1 / (2*pi*\<i>) *
+                                 integral {0..x} (\<lambda>t. 1/(\<gamma> t - z) * vector_derivative \<gamma> (at t))"])
+    apply (rule continuous_intros)+
+    apply (rule indefinite_integral_continuous)
+    apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on])
+      using assms
+    apply (simp add: *)
+    done
+qed
+
+lemma winding_number_ivt_pos:
+    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> Re(winding_number \<gamma> z)"
+      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
+  apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
+  apply (simp add:)
+  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
+  using assms
+  apply (auto simp: path_image_def image_def)
+  done
+
+lemma winding_number_ivt_neg:
+    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "Re(winding_number \<gamma> z) \<le> w" "w \<le> 0"
+      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
+  apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
+  apply (simp add:)
+  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
+  using assms
+  apply (auto simp: path_image_def image_def)
+  done
+
+lemma winding_number_ivt_abs:
+    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>"
+      shows "\<exists>t \<in> {0..1}. \<bar>Re (winding_number (subpath 0 t \<gamma>) z)\<bar> = w"
+  using assms winding_number_ivt_pos [of \<gamma> z w] winding_number_ivt_neg [of \<gamma> z "-w"]
+  by force
+
+lemma winding_number_lt_half_lemma:
+  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}"
+    shows "Re(winding_number \<gamma> z) < 1/2"
+proof -
+  { assume "Re(winding_number \<gamma> z) \<ge> 1/2"
+    then obtain t::real where t: "0 \<le> t" "t \<le> 1" and sub12: "Re (winding_number (subpath 0 t \<gamma>) z) = 1/2"
+      using winding_number_ivt_pos [OF \<gamma> z, of "1/2"] by auto
+    have gt: "\<gamma> t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \<gamma>) z)))) * (\<gamma> 0 - z))"
+      using winding_number_exp_2pi [of "subpath 0 t \<gamma>" z]
+      apply (simp add: t \<gamma> valid_path_imp_path)
+      using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12)
+    have "b < a \<bullet> \<gamma> 0"
+    proof -
+      have "\<gamma> 0 \<in> {c. b < a \<bullet> c}"
+        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one)
+      thus ?thesis
+        by blast
+    qed
+    moreover have "b < a \<bullet> \<gamma> t"
+    proof -
+      have "\<gamma> t \<in> {c. b < a \<bullet> c}"
+        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t)
+      thus ?thesis
+        by blast
+    qed
+    ultimately have "0 < a \<bullet> (\<gamma> 0 - z)" "0 < a \<bullet> (\<gamma> t - z)" using az
+      by (simp add: inner_diff_right)+
+    then have False
+      by (simp add: gt inner_mult_right mult_less_0_iff)
+  }
+  then show ?thesis by force
+qed
+
+lemma winding_number_lt_half:
+  assumes "valid_path \<gamma>" "a \<bullet> z \<le> b" "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
+    shows "\<bar>Re (winding_number \<gamma> z)\<bar> < 1/2"
+proof -
+  have "z \<notin> path_image \<gamma>" using assms by auto
+  with assms show ?thesis
+    apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1)
+    apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \<gamma> z a b]
+                 winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath)
+    done
+qed
+
+lemma winding_number_le_half:
+  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
+      and anz: "a \<noteq> 0" and azb: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w \<ge> b}"
+    shows "\<bar>Re (winding_number \<gamma> z)\<bar> \<le> 1/2"
+proof -
+  { assume wnz_12: "\<bar>Re (winding_number \<gamma> z)\<bar> > 1/2"
+    have "isCont (winding_number \<gamma>) z"
+      by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z)
+    then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < \<bar>Re(winding_number \<gamma> z)\<bar> - 1/2"
+      using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast
+    define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a"
+    have *: "a \<bullet> z' \<le> b - d / 3 * cmod a"
+      unfolding z'_def inner_mult_right' divide_inverse
+      apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz)
+      apply (metis \<open>0 < d\<close> add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral)
+      done
+    have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2"
+      using d [of z'] anz \<open>d>0\<close> by (simp add: dist_norm z'_def)
+    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)"
+      by simp
+    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>"
+      using abs_Re_le_cmod [of "winding_number \<gamma> z' - winding_number \<gamma> z"] by simp
+    then have wnz_12': "\<bar>Re (winding_number \<gamma> z')\<bar> > 1/2"
+      by linarith
+    moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2"
+      apply (rule winding_number_lt_half [OF \<gamma> *])
+      using azb \<open>d>0\<close> pag
+      apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD)
+      done
+    ultimately have False
+      by simp
+  }
+  then show ?thesis by force
+qed
+
+lemma winding_number_lt_half_linepath: "z \<notin> closed_segment a b \<Longrightarrow> \<bar>Re (winding_number (linepath a b) z)\<bar> < 1/2"
+  using separating_hyperplane_closed_point [of "closed_segment a b" z]
+  apply auto
+  apply (simp add: closed_segment_def)
+  apply (drule less_imp_le)
+  apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]])
+  apply (auto simp: segment)
+  done
+
+
+text\<open> Positivity of WN for a linepath.\<close>
+lemma winding_number_linepath_pos_lt:
+    assumes "0 < Im ((b - a) * cnj (b - z))"
+      shows "0 < Re(winding_number(linepath a b) z)"
+proof -
+  have z: "z \<notin> path_image (linepath a b)"
+    using assms
+    by (simp add: closed_segment_def) (force simp: algebra_simps)
+  show ?thesis
+    apply (rule winding_number_pos_lt [OF valid_path_linepath z assms])
+    apply (simp add: linepath_def algebra_simps)
+    done
+qed
+
+
+subsection\<open>Cauchy's integral formula, again for a convex enclosing set.\<close>
+
+lemma Cauchy_integral_formula_weak:
+    assumes s: "convex s" and "finite k" and conf: "continuous_on s f"
+        and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
+        and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>"
+        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
+proof -
+  obtain f' where f': "(f has_field_derivative f') (at z)"
+    using fcd [OF z] by (auto simp: field_differentiable_def)
+  have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
+  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
+  proof (cases "x = z")
+    case True then show ?thesis
+      apply (simp add: continuous_within)
+      apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"])
+      using has_field_derivative_at_within DERIV_within_iff f'
+      apply (fastforce simp add:)+
+      done
+  next
+    case False
+    then have dxz: "dist x z > 0" by auto
+    have cf: "continuous (at x within s) f"
+      using conf continuous_on_eq_continuous_within that by blast
+    have "continuous (at x within s) (\<lambda>w. (f w - f z) / (w - z))"
+      by (rule cf continuous_intros | simp add: False)+
+    then show ?thesis
+      apply (rule continuous_transform_within [OF _ dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"])
+      apply (force simp: dist_commute)
+      done
+  qed
+  have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
+  have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
+    apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop])
+    using c apply (force simp: continuous_on_eq_continuous_within)
+    apply (rename_tac w)
+    apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in field_differentiable_transform_within)
+    apply (simp_all add: dist_pos_lt dist_commute)
+    apply (metis less_irrefl)
+    apply (rule derivative_intros fcd | simp)+
+    done
+  show ?thesis
+    apply (rule has_contour_integral_eq)
+    using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
+    apply (auto simp: mult_ac divide_simps)
+    done
+qed
+
+theorem Cauchy_integral_formula_convex_simple:
+    "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z};
+      pathfinish \<gamma> = pathstart \<gamma>\<rbrakk>
+     \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
+  apply (rule Cauchy_integral_formula_weak [where k = "{}"])
+  using holomorphic_on_imp_continuous_on
+  by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE)
+
+
+subsection\<open>Homotopy forms of Cauchy's theorem\<close>
+
+proposition Cauchy_theorem_homotopic:
+    assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h"
+        and "open s" and f: "f holomorphic_on s"
+        and vpg: "valid_path g" and vph: "valid_path h"
+    shows "contour_integral g f = contour_integral h f"
+proof -
+  have pathsf: "linked_paths atends g h"
+    using hom  by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop)
+  obtain k :: "real \<times> real \<Rightarrow> complex"
+    where contk: "continuous_on ({0..1} \<times> {0..1}) k"
+      and ks: "k ` ({0..1} \<times> {0..1}) \<subseteq> s"
+      and k [simp]: "\<forall>x. k (0, x) = g x" "\<forall>x. k (1, x) = h x"
+      and ksf: "\<forall>t\<in>{0..1}. linked_paths atends g (\<lambda>x. k (t, x))"
+      using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm)
+  have ucontk: "uniformly_continuous_on ({0..1} \<times> {0..1}) k"
+    by (blast intro: compact_Times compact_uniformly_continuous [OF contk])
+  { fix t::real assume t: "t \<in> {0..1}"
+    have pak: "path (k o (\<lambda>u. (t, u)))"
+      unfolding path_def
+      apply (rule continuous_intros continuous_on_subset [OF contk])+
+      using t by force
+    have pik: "path_image (k \<circ> Pair t) \<subseteq> s"
+      using ks t by (auto simp: path_image_def)
+    obtain e where "e>0" and e:
+         "\<And>g h. \<lbrakk>valid_path g; valid_path h;
+                  \<forall>u\<in>{0..1}. cmod (g u - (k \<circ> Pair t) u) < e \<and> cmod (h u - (k \<circ> Pair t) u) < e;
+                  linked_paths atends g h\<rbrakk>
+                 \<Longrightarrow> contour_integral h f = contour_integral g f"
+      using contour_integral_nearby [OF \<open>open s\<close> pak pik, of atends] f by metis
+    obtain d where "d>0" and d:
+        "\<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"
+      by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \<open>e>0\<close>)
+    { fix t1 t2
+      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"
+      have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e"
+        using \<open>e > 0\<close>
+        apply (rule_tac y = k1 in norm_triangle_half_l)
+        apply (auto simp: norm_minus_commute intro: order_less_trans)
+        done
+      have "\<exists>d>0. \<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
+                          (\<forall>u\<in>{0..1}. cmod (g1 u - k (t1, u)) < d \<and> cmod (g2 u - k (t2, u)) < d) \<and>
+                          linked_paths atends g1 g2 \<longrightarrow>
+                          contour_integral g2 f = contour_integral g1 f"
+        apply (rule_tac x="e/4" in exI)
+        using t t1 t2 ltd \<open>e > 0\<close>
+        apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1)
+        done
+    }
+    then have "\<exists>e. 0 < e \<and>
+              (\<forall>t1 t2. t1 \<in> {0..1} \<and> t2 \<in> {0..1} \<and> \<bar>t1 - t\<bar> < e \<and> \<bar>t2 - t\<bar> < e
+                \<longrightarrow> (\<exists>d. 0 < d \<and>
+                     (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
+                       (\<forall>u \<in> {0..1}.
+                          norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
+                          linked_paths atends g1 g2
+                          \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
+      by (rule_tac x=d in exI) (simp add: \<open>d > 0\<close>)
+  }
+  then obtain ee where ee:
+       "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0 \<and>
+          (\<forall>t1 t2. t1 \<in> {0..1} \<longrightarrow> t2 \<in> {0..1} \<longrightarrow> \<bar>t1 - t\<bar> < ee t \<longrightarrow> \<bar>t2 - t\<bar> < ee t
+            \<longrightarrow> (\<exists>d. 0 < d \<and>
+                 (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
+                   (\<forall>u \<in> {0..1}.
+                      norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
+                      linked_paths atends g1 g2
+                      \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
+    by metis
+  note ee_rule = ee [THEN conjunct2, rule_format]
+  define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}"
+  have "\<forall>t \<in> C. open t" by (simp add: C_def)
+  moreover have "{0..1} \<subseteq> \<Union>C"
+    using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
+  ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
+    by (rule compactE [OF compact_interval])
+  define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
+  have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
+  define e where "e = Min (ee ` kk)"
+  have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk"
+    using C' by (auto simp: kk_def C_def)
+  have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0"
+    by (simp add: kk_def ee)
+  moreover have "finite kk"
+    using \<open>finite C'\<close> kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD)
+  moreover have "kk \<noteq> {}" using \<open>{0..1} \<subseteq> \<Union>C'\<close> C'_eq by force
+  ultimately have "e > 0"
+    using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def)
+  then obtain N::nat where "N > 0" and N: "1/N < e/3"
+    by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral)
+  have e_le_ee: "\<And>i. i \<in> kk \<Longrightarrow> e \<le> ee i"
+    using \<open>finite kk\<close> by (simp add: e_def Min_le_iff [of "ee ` kk"])
+  have plus: "\<exists>t \<in> kk. x \<in> ball t (ee t / 3)" if "x \<in> {0..1}" for x
+    using C' subsetD [OF C'01 that]  unfolding C'_eq by blast
+  have [OF order_refl]:
+      "\<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
+                        \<longrightarrow> contour_integral j f = contour_integral g f)"
+       if "n \<le> N" for n
+  using that
+  proof (induct n)
+    case 0 show ?case using ee_rule [of 0 0 0]
+      apply clarsimp
+      apply (rule_tac x=d in exI, safe)
+      by (metis diff_self vpg norm_zero)
+  next
+    case (Suc n)
+    then have N01: "n/N \<in> {0..1}" "(Suc n)/N \<in> {0..1}"  by auto
+    then obtain t where t: "t \<in> kk" "n/N \<in> ball t (ee t / 3)"
+      using plus [of "n/N"] by blast
+    then have nN_less: "\<bar>n/N - t\<bar> < ee t"
+      by (simp add: dist_norm del: less_divide_eq_numeral1)
+    have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t"
+      using t N \<open>N > 0\<close> e_le_ee [of t]
+      by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps)
+    have t01: "t \<in> {0..1}" using \<open>kk \<subseteq> {0..1}\<close> \<open>t \<in> kk\<close> by blast
+    obtain d1 where "d1 > 0" and d1:
+        "\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2;
+                   \<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1;
+                   linked_paths atends g1 g2\<rbrakk>
+                   \<Longrightarrow> contour_integral g2 f = contour_integral g1 f"
+      using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce
+    have "n \<le> N" using Suc.prems by auto
+    with Suc.hyps
+    obtain d2 where "d2 > 0"
+      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>
+                     \<Longrightarrow> contour_integral j f = contour_integral g f"
+        by auto
+    have "continuous_on {0..1} (k o (\<lambda>u. (n/N, u)))"
+      apply (rule continuous_intros continuous_on_subset [OF contk])+
+      using N01 by auto
+    then have pkn: "path (\<lambda>u. k (n/N, u))"
+      by (simp add: path_def)
+    have min12: "min d1 d2 > 0" by (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>)
+    obtain p where "polynomial_function p"
+        and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))"
+                 "pathfinish p = pathfinish (\<lambda>u. k (n/N, u))"
+        and pk_le:  "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (p t - k (n/N, t)) < min d1 d2"
+      using path_approx_polynomial_function [OF pkn min12] by blast
+    then have vpp: "valid_path p" using valid_path_polynomial_function by blast
+    have lpa: "linked_paths atends g p"
+      by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf)
+    show ?case
+      apply (rule_tac x="min d1 d2" in exI)
+      apply (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>, clarify)
+      apply (rule_tac s="contour_integral p f" in trans)
+      using pk_le N01(1) ksf pathfinish_def pathstart_def
+      apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf)
+      using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf)
+      done
+  qed
+  then obtain d where "0 < d"
+                       "\<And>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (1,u)) < d) \<and>
+                            linked_paths atends g j
+                            \<Longrightarrow> contour_integral j f = contour_integral g f"
+    using \<open>N>0\<close> by auto
+  then have "linked_paths atends g h \<Longrightarrow> contour_integral h f = contour_integral g f"
+    using \<open>N>0\<close> vph by fastforce
+  then show ?thesis
+    by (simp add: pathsf)
+qed
+
+proposition Cauchy_theorem_homotopic_paths:
+    assumes hom: "homotopic_paths s g h"
+        and "open s" and f: "f holomorphic_on s"
+        and vpg: "valid_path g" and vph: "valid_path h"
+    shows "contour_integral g f = contour_integral h f"
+  using Cauchy_theorem_homotopic [of True s g h] assms by simp
+
+proposition Cauchy_theorem_homotopic_loops:
+    assumes hom: "homotopic_loops s g h"
+        and "open s" and f: "f holomorphic_on s"
+        and vpg: "valid_path g" and vph: "valid_path h"
+    shows "contour_integral g f = contour_integral h f"
+  using Cauchy_theorem_homotopic [of False s g h] assms by simp
+
+lemma has_contour_integral_newpath:
+    "\<lbrakk>(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\<rbrakk>
+     \<Longrightarrow> (f has_contour_integral y) g"
+  using has_contour_integral_integral contour_integral_unique by auto
+
+lemma Cauchy_theorem_null_homotopic:
+     "\<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"
+  apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp)
+  using contour_integrable_holomorphic_simple
+    apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset)
+  by (simp add: Cauchy_theorem_homotopic_loops)
+
+
+
+subsection\<open>More winding number properties\<close>
+
+text\<open>including the fact that it's +-1 inside a simple closed curve.\<close>
+
+lemma winding_number_homotopic_paths:
+    assumes "homotopic_paths (-{z}) g h"
+      shows "winding_number g z = winding_number h z"
+proof -
+  have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto
+  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
+    using homotopic_paths_imp_subset [OF assms] by auto
+  ultimately obtain d e where "d > 0" "e > 0"
+      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>
+            \<Longrightarrow> homotopic_paths (-{z}) g p"
+      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>
+            \<Longrightarrow> homotopic_paths (-{z}) h q"
+    using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force
+  obtain p where p:
+       "valid_path p" "z \<notin> path_image p"
+       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
+       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
+       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
+    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
+  obtain q where q:
+       "valid_path q" "z \<notin> path_image q"
+       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
+       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
+       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
+    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
+  have gp: "homotopic_paths (- {z}) g p"
+    by (simp add: d p valid_path_imp_path norm_minus_commute gp_less)
+  have hq: "homotopic_paths (- {z}) h q"
+    by (simp add: e q valid_path_imp_path norm_minus_commute hq_less)
+  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
+    apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"])
+    apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms)
+    apply (auto intro!: holomorphic_intros simp: p q)
+    done
+  then show ?thesis
+    by (simp add: pap paq)
+qed
+
+lemma winding_number_homotopic_loops:
+    assumes "homotopic_loops (-{z}) g h"
+      shows "winding_number g z = winding_number h z"
+proof -
+  have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto
+  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
+    using homotopic_loops_imp_subset [OF assms] by auto
+  moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h"
+    using homotopic_loops_imp_loop [OF assms] by auto
+  ultimately obtain d e where "d > 0" "e > 0"
+      and d: "\<And>p. \<lbrakk>path p; pathfinish p = pathstart p; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
+            \<Longrightarrow> homotopic_loops (-{z}) g p"
+      and e: "\<And>q. \<lbrakk>path q; pathfinish q = pathstart q; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
+            \<Longrightarrow> homotopic_loops (-{z}) h q"
+    using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force
+  obtain p where p:
+       "valid_path p" "z \<notin> path_image p"
+       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
+       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
+       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
+    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
+  obtain q where q:
+       "valid_path q" "z \<notin> path_image q"
+       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
+       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
+       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
+    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
+  have gp: "homotopic_loops (- {z}) g p"
+    by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path)
+  have hq: "homotopic_loops (- {z}) h q"
+    by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path)
+  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
+    apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"])
+    apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms)
+    apply (auto intro!: holomorphic_intros simp: p q)
+    done
+  then show ?thesis
+    by (simp add: pap paq)
+qed
+
+lemma winding_number_paths_linear_eq:
+  "\<lbrakk>path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g;
+    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
+        \<Longrightarrow> winding_number h z = winding_number g z"
+  by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: )
+
+lemma winding_number_loops_linear_eq:
+  "\<lbrakk>path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h;
+    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
+        \<Longrightarrow> winding_number h z = winding_number g z"
+  by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: )
+
+lemma winding_number_nearby_paths_eq:
+     "\<lbrakk>path g; path h;
+      pathstart h = pathstart g; pathfinish h = pathfinish g;
+      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
+      \<Longrightarrow> winding_number h z = winding_number g z"
+  by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq)
+
+lemma winding_number_nearby_loops_eq:
+     "\<lbrakk>path g; path h;
+      pathfinish g = pathstart g;
+        pathfinish h = pathstart h;
+      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
+      \<Longrightarrow> winding_number h z = winding_number g z"
+  by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq)
+
+
+proposition winding_number_subpath_combine:
+    "\<lbrakk>path g; z \<notin> path_image g;
+      u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
+      \<Longrightarrow> winding_number (subpath u v g) z + winding_number (subpath v w g) z =
+          winding_number (subpath u w g) z"
+apply (rule trans [OF winding_number_join [THEN sym]
+                      winding_number_homotopic_paths [OF homotopic_join_subpaths]])
+apply (auto dest: path_image_subpath_subset)
+done
+
+
+subsection\<open>Partial circle path\<close>
+
+definition part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
+  where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))"
+
+lemma pathstart_part_circlepath [simp]:
+     "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)"
+by (metis part_circlepath_def pathstart_def pathstart_linepath)
+
+lemma pathfinish_part_circlepath [simp]:
+     "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)"
+by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
+
+proposition has_vector_derivative_part_circlepath [derivative_intros]:
+    "((part_circlepath z r s t) has_vector_derivative
+      (\<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)))
+     (at x within X)"
+  apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real)
+  apply (rule has_vector_derivative_real_complex)
+  apply (rule derivative_eq_intros | simp)+
+  done
+
+corollary vector_derivative_part_circlepath:
+    "vector_derivative (part_circlepath z r s t) (at x) =
+       \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
+  using has_vector_derivative_part_circlepath vector_derivative_at by blast
+
+corollary vector_derivative_part_circlepath01:
+    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
+     \<Longrightarrow> vector_derivative (part_circlepath z r s t) (at x within {0..1}) =
+          \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
+  using has_vector_derivative_part_circlepath
+  by (auto simp: vector_derivative_at_within_ivl)
+
+lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)"
+  apply (simp add: valid_path_def)
+  apply (rule C1_differentiable_imp_piecewise)
+  apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath
+              intro!: continuous_intros)
+  done
+
+lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)"
+  by (simp add: valid_path_imp_path)
+
+proposition path_image_part_circlepath:
+  assumes "s \<le> t"
+    shows "path_image (part_circlepath z r s t) = {z + r * exp(\<i> * of_real x) | x. s \<le> x \<and> x \<le> t}"
+proof -
+  { fix z::real
+    assume "0 \<le> z" "z \<le> 1"
+    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"
+      apply (rule_tac x="(1 - z) * s + z * t" in exI)
+      apply (simp add: linepath_def scaleR_conv_of_real algebra_simps)
+      apply (rule conjI)
+      using mult_right_mono apply blast
+      using affine_ineq  by (metis "mult.commute")
+  }
+  moreover
+  { fix z
+    assume "s \<le> z" "z \<le> t"
+    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}"
+      apply (rule_tac x="(z - s)/(t - s)" in image_eqI)
+      apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq)
+      apply (auto simp: algebra_simps divide_simps)
+      done
+  }
+  ultimately show ?thesis
+    by (fastforce simp add: path_image_def part_circlepath_def)
+qed
+
+corollary path_image_part_circlepath_subset:
+    "\<lbrakk>s \<le> t; 0 \<le> r\<rbrakk> \<Longrightarrow> path_image(part_circlepath z r s t) \<subseteq> sphere z r"
+by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult)
+
+proposition in_path_image_part_circlepath:
+  assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r"
+    shows "norm(w - z) = r"
+proof -
+  have "w \<in> {c. dist z c = r}"
+    by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms)
+  thus ?thesis
+    by (simp add: dist_norm norm_minus_commute)
+qed
+
+proposition finite_bounded_log: "finite {z::complex. norm z \<le> b \<and> exp z = w}"
+proof (cases "w = 0")
+  case True then show ?thesis by auto
+next
+  case False
+  have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \<i>) \<le> b + cmod (Ln w)}"
+    apply (simp add: norm_mult finite_int_iff_bounded_le)
+    apply (rule_tac x="\<lfloor>(b + cmod (Ln w)) / (2*pi)\<rfloor>" in exI)
+    apply (auto simp: divide_simps le_floor_iff)
+    done
+  have [simp]: "\<And>P f. {z. P z \<and> (\<exists>n. z = f n)} = f ` {n. P (f n)}"
+    by blast
+  show ?thesis
+    apply (subst exp_Ln [OF False, symmetric])
+    apply (simp add: exp_eq)
+    using norm_add_leD apply (fastforce intro: finite_subset [OF _ *])
+    done
+qed
+
+lemma finite_bounded_log2:
+  fixes a::complex
+    assumes "a \<noteq> 0"
+    shows "finite {z. norm z \<le> b \<and> exp(a*z) = w}"
+proof -
+  have *: "finite ((\<lambda>z. z / a) ` {z. cmod z \<le> b * cmod a \<and> exp z = w})"
+    by (rule finite_imageI [OF finite_bounded_log])
+  show ?thesis
+    by (rule finite_subset [OF _ *]) (force simp: assms norm_mult)
+qed
+
+proposition has_contour_integral_bound_part_circlepath_strong:
+  assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)"
+      and "finite k" and le: "0 \<le> B" "0 < r" "s \<le> t"
+      and B: "\<And>x. x \<in> path_image(part_circlepath z r s t) - k \<Longrightarrow> norm(f x) \<le> B"
+    shows "cmod i \<le> B * r * (t - s)"
+proof -
+  consider "s = t" | "s < t" using \<open>s \<le> t\<close> by linarith
+  then show ?thesis
+  proof cases
+    case 1 with fi [unfolded has_contour_integral]
+    have "i = 0"  by (simp add: vector_derivative_part_circlepath)
+    with assms show ?thesis by simp
+  next
+    case 2
+    have [simp]: "\<bar>r\<bar> = r" using \<open>r > 0\<close> by linarith
+    have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s"
+      by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff)
+    have "finite (part_circlepath z r s t -` {y} \<inter> {0..1})" if "y \<in> k" for y
+    proof -
+      define w where "w = (y - z)/of_real r / exp(\<i> * of_real s)"
+      have fin: "finite (of_real -` {z. cmod z \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})"
+        apply (rule finite_vimageI [OF finite_bounded_log2])
+        using \<open>s < t\<close> apply (auto simp: inj_of_real)
+        done
+      show ?thesis
+        apply (simp add: part_circlepath_def linepath_def vimage_def)
+        apply (rule finite_subset [OF _ fin])
+        using le
+        apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff)
+        done
+    qed
+    then have fin01: "finite ((part_circlepath z r s t) -` k \<inter> {0..1})"
+      by (rule finite_finite_vimage_IntI [OF \<open>finite k\<close>])
+    have **: "((\<lambda>x. if (part_circlepath z r s t x) \<in> k then 0
+                    else f(part_circlepath z r s t x) *
+                       vector_derivative (part_circlepath z r s t) (at x)) has_integral i)  {0..1}"
+      apply (rule has_integral_spike
+              [where f = "\<lambda>x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"])
+      apply (rule negligible_finite [OF fin01])
+      using fi has_contour_integral
+      apply auto
+      done
+    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"
+      by (auto intro!: B [unfolded path_image_def image_def, simplified])
+    show ?thesis
+      apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified])
+      using assms apply force
+      apply (simp add: norm_mult vector_derivative_part_circlepath)
+      using le * "2" \<open>r > 0\<close> by auto
+  qed
+qed
+
+corollary has_contour_integral_bound_part_circlepath:
+      "\<lbrakk>(f has_contour_integral i) (part_circlepath z r s t);
+        0 \<le> B; 0 < r; s \<le> t;
+        \<And>x. x \<in> path_image(part_circlepath z r s t) \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
+       \<Longrightarrow> norm i \<le> B*r*(t - s)"
+  by (auto intro: has_contour_integral_bound_part_circlepath_strong)
+
+proposition contour_integrable_continuous_part_circlepath:
+     "continuous_on (path_image (part_circlepath z r s t)) f
+      \<Longrightarrow> f contour_integrable_on (part_circlepath z r s t)"
+  apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def)
+  apply (rule integrable_continuous_real)
+  apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl])
+  done
+
+proposition winding_number_part_circlepath_pos_less:
+  assumes "s < t" and no: "norm(w - z) < r"
+    shows "0 < Re (winding_number(part_circlepath z r s t) w)"
+proof -
+  have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2)
+  note valid_path_part_circlepath
+  moreover have " w \<notin> path_image (part_circlepath z r s t)"
+    using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def)
+  moreover have "0 < r * (t - s) * (r - cmod (w - z))"
+    using assms by (metis \<open>0 < r\<close> diff_gt_0_iff_gt mult_pos_pos)
+  ultimately show ?thesis
+    apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"])
+    apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac)
+    apply (rule mult_left_mono)+
+    using Re_Im_le_cmod [of "w-z" "linepath s t x" for x]
+    apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square])
+    using assms \<open>0 < r\<close> by auto
+qed
+
+proposition simple_path_part_circlepath:
+    "simple_path(part_circlepath z r s t) \<longleftrightarrow> (r \<noteq> 0 \<and> s \<noteq> t \<and> \<bar>s - t\<bar> \<le> 2*pi)"
+proof (cases "r = 0 \<or> s = t")
+  case True
+  then show ?thesis
+    apply (rule disjE)
+    apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+
+    done
+next
+  case False then have "r \<noteq> 0" "s \<noteq> t" by auto
+  have *: "\<And>x y z s t. \<i>*((1 - x) * s + x * t) = \<i>*(((1 - y) * s + y * t)) + z  \<longleftrightarrow> \<i>*(x - y) * (t - s) = z"
+    by (simp add: algebra_simps)
+  have abs01: "\<And>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1
+                      \<Longrightarrow> (x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0 \<longleftrightarrow> \<bar>x - y\<bar> \<in> {0,1})"
+    by auto
+  have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
+    by force
+  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>
+                  (\<exists>n. \<bar>x - y\<bar> * (t - s) = 2 * (of_int n * pi))"
+    by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real]
+                    intro: exI [where x = "-n" for n])
+  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"
+    apply (rule ccontr)
+    apply (drule_tac x="2*pi / \<bar>t - s\<bar>" in spec)
+    using False
+    apply (simp add: abs_minus_commute divide_simps)
+    apply (frule_tac x=1 in spec)
+    apply (drule_tac x="-1" in spec)
+    apply (simp add:)
+    done
+  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
+  proof -
+    have "t-s = 2 * (real_of_int n * pi)/x"
+      using that by (simp add: field_simps)
+    then show ?thesis by (metis abs_minus_commute)
+  qed
+  show ?thesis using False
+    apply (simp add: simple_path_def path_part_circlepath)
+    apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01  del: Set.insert_iff)
+    apply (subst abs_away)
+    apply (auto simp: 1)
+    apply (rule ccontr)
+    apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD)
+    done
+qed
+
+proposition arc_part_circlepath:
+  assumes "r \<noteq> 0" "s \<noteq> t" "\<bar>s - t\<bar> < 2*pi"
+    shows "arc (part_circlepath z r s t)"
+proof -
+  have *: "x = y" if eq: "\<i> * (linepath s t x) = \<i> * (linepath s t y) + 2 * of_int n * complex_of_real pi * \<i>"
+                  and x: "x \<in> {0..1}" and y: "y \<in> {0..1}" for x y n
+    proof -
+      have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi"
+        by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq)
+      then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))"
+        by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re])
+      then have st: "x \<noteq> y \<Longrightarrow> (s-t) = (of_int n * (pi * 2) / (y-x))"
+        by (force simp: field_simps)
+      show ?thesis
+        apply (rule ccontr)
+        using assms x y
+        apply (simp add: st abs_mult field_simps)
+        using st
+        apply (auto simp: dest: of_int_lessD)
+        done
+    qed
+  show ?thesis
+    using assms
+    apply (simp add: arc_def)
+    apply (simp add: part_circlepath_def inj_on_def exp_eq)
+    apply (blast intro: *)
+    done
+qed
+
+
+subsection\<open>Special case of one complete circle\<close>
+
+definition circlepath :: "[complex, real, real] \<Rightarrow> complex"
+  where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)"
+
+lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * \<i> * of_real x))"
+  by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps)
+
+lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r"
+  by (simp add: circlepath_def)
+
+lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r"
+  by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute)
+
+lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)"
+proof -
+  have "z + of_real r * exp (2 * pi * \<i> * (x + 1 / 2)) =
+        z + of_real r * exp (2 * pi * \<i> * x + pi * \<i>)"
+    by (simp add: divide_simps) (simp add: algebra_simps)
+  also have "... = z - r * exp (2 * pi * \<i> * x)"
+    by (simp add: exp_add)
+  finally show ?thesis
+    by (simp add: circlepath path_image_def sphere_def dist_norm)
+qed
+
+lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x"
+  using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x]
+  by (simp add: add.commute)
+
+lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)"
+  using circlepath_add1 [of z r "x-1/2"]
+  by (simp add: add.commute)
+
+lemma path_image_circlepath_minus_subset:
+     "path_image (circlepath z (-r)) \<subseteq> path_image (circlepath z r)"
+  apply (simp add: path_image_def image_def circlepath_minus, clarify)
+  apply (case_tac "xa \<le> 1/2", force)
+  apply (force simp add: circlepath_add_half)+
+  done
+
+lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)"
+  using path_image_circlepath_minus_subset by fastforce
+
+proposition has_vector_derivative_circlepath [derivative_intros]:
+ "((circlepath z r) has_vector_derivative (2 * pi * \<i> * r * exp (2 * of_real pi * \<i> * of_real x)))
+   (at x within X)"
+  apply (simp add: circlepath_def scaleR_conv_of_real)
+  apply (rule derivative_eq_intros)
+  apply (simp add: algebra_simps)
+  done
+
+corollary vector_derivative_circlepath:
+   "vector_derivative (circlepath z r) (at x) =
+    2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
+using has_vector_derivative_circlepath vector_derivative_at by blast
+
+corollary vector_derivative_circlepath01:
+    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
+     \<Longrightarrow> vector_derivative (circlepath z r) (at x within {0..1}) =
+          2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
+  using has_vector_derivative_circlepath
+  by (auto simp: vector_derivative_at_within_ivl)
+
+lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)"
+  by (simp add: circlepath_def)
+
+lemma path_circlepath [simp]: "path (circlepath z r)"
+  by (simp add: valid_path_imp_path)
+
+lemma path_image_circlepath_nonneg:
+  assumes "0 \<le> r" shows "path_image (circlepath z r) = sphere z r"
+proof -
+  have *: "x \<in> (\<lambda>u. z + (cmod (x - z)) * exp (\<i> * (of_real u * (of_real pi * 2)))) ` {0..1}" for x
+  proof (cases "x = z")
+    case True then show ?thesis by force
+  next
+    case False
+    define w where "w = x - z"
+    then have "w \<noteq> 0" by (simp add: False)
+    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)"
+      using cis_conv_exp complex_eq_iff by auto
+    show ?thesis
+      apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"])
+      apply (simp add: divide_simps \<open>w \<noteq> 0\<close> cmod_power2 [symmetric])
+      apply (rule_tac x="t / (2*pi)" in image_eqI)
+      apply (simp add: divide_simps \<open>w \<noteq> 0\<close>)
+      using False **
+      apply (auto simp: w_def)
+      done
+  qed
+  show ?thesis
+    unfolding circlepath path_image_def sphere_def dist_norm
+    by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *)
+qed
+
+proposition path_image_circlepath [simp]:
+    "path_image (circlepath z r) = sphere z \<bar>r\<bar>"
+  using path_image_circlepath_minus
+  by (force simp add: path_image_circlepath_nonneg abs_if)
+
+lemma has_contour_integral_bound_circlepath_strong:
+      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
+        finite k; 0 \<le> B; 0 < r;
+        \<And>x. \<lbrakk>norm(x - z) = r; x \<notin> k\<rbrakk> \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
+        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
+  unfolding circlepath_def
+  by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong)
+
+corollary has_contour_integral_bound_circlepath:
+      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
+        0 \<le> B; 0 < r; \<And>x. norm(x - z) = r \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
+        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
+  by (auto intro: has_contour_integral_bound_circlepath_strong)
+
+proposition contour_integrable_continuous_circlepath:
+    "continuous_on (path_image (circlepath z r)) f
+     \<Longrightarrow> f contour_integrable_on (circlepath z r)"
+  by (simp add: circlepath_def contour_integrable_continuous_part_circlepath)
+
+lemma simple_path_circlepath: "simple_path(circlepath z r) \<longleftrightarrow> (r \<noteq> 0)"
+  by (simp add: circlepath_def simple_path_part_circlepath)
+
+lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \<Longrightarrow> w \<notin> path_image (circlepath z r)"
+  by (simp add: sphere_def dist_norm norm_minus_commute)
+
+proposition contour_integral_circlepath:
+     "0 < r \<Longrightarrow> contour_integral (circlepath z r) (\<lambda>w. 1 / (w - z)) = 2 * complex_of_real pi * \<i>"
+  apply (rule contour_integral_unique)
+  apply (simp add: has_contour_integral_def)
+  apply (subst has_integral_cong)
+  apply (simp add: vector_derivative_circlepath01)
+  using has_integral_const_real [of _ 0 1]
+  apply (force simp: circlepath)
+  done
+
+lemma winding_number_circlepath_centre: "0 < r \<Longrightarrow> winding_number (circlepath z r) z = 1"
+  apply (rule winding_number_unique_loop)
+  apply (simp_all add: sphere_def valid_path_imp_path)
+  apply (rule_tac x="circlepath z r" in exI)
+  apply (simp add: sphere_def contour_integral_circlepath)
+  done
+
+proposition winding_number_circlepath:
+  assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1"
+proof (cases "w = z")
+  case True then show ?thesis
+    using assms winding_number_circlepath_centre by auto
+next
+  case False
+  have [simp]: "r > 0"
+    using assms le_less_trans norm_ge_zero by blast
+  define r' where "r' = norm(w - z)"
+  have "r' < r"
+    by (simp add: assms r'_def)
+  have disjo: "cball z r' \<inter> sphere z r = {}"
+    using \<open>r' < r\<close> by (force simp: cball_def sphere_def)
+  have "winding_number(circlepath z r) w = winding_number(circlepath z r) z"
+    apply (rule winding_number_around_inside [where s = "cball z r'"])
+    apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre)
+    apply (simp_all add: False r'_def dist_norm norm_minus_commute)
+    done
+  also have "... = 1"
+    by (simp add: winding_number_circlepath_centre)
+  finally show ?thesis .
+qed
+
+
+text\<open> Hence the Cauchy formula for points inside a circle.\<close>
+
+theorem Cauchy_integral_circlepath:
+  assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r"
+  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
+         (circlepath z r)"
+proof -
+  have "r > 0"
+    using assms le_less_trans norm_ge_zero by blast
+  have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
+        (circlepath z r)"
+    apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"])
+    using assms  \<open>r > 0\<close>
+    apply (simp_all add: dist_norm norm_minus_commute)
+    apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute)
+    apply (simp add: cball_def sphere_def dist_norm, clarify)
+    apply (simp add:)
+    by (metis dist_commute dist_norm less_irrefl)
+  then show ?thesis
+    by (simp add: winding_number_circlepath assms)
+qed
+
+corollary Cauchy_integral_circlepath_simple:
+  assumes "f holomorphic_on cball z r" "norm(w - z) < r"
+  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
+         (circlepath z r)"
+using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath)
+
+
+lemma no_bounded_connected_component_imp_winding_number_zero:
+  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
+      and nb: "\<And>z. bounded (connected_component_set (- s) z) \<longrightarrow> z \<in> s"
+  shows "winding_number g z = 0"
+apply (rule winding_number_zero_in_outside)
+apply (simp_all add: assms)
+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)
+
+lemma no_bounded_path_component_imp_winding_number_zero:
+  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
+      and nb: "\<And>z. bounded (path_component_set (- s) z) \<longrightarrow> z \<in> s"
+  shows "winding_number g z = 0"
+apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g])
+by (simp add: bounded_subset nb path_component_subset_connected_component)
+
+
+subsection\<open> Uniform convergence of path integral\<close>
+
+text\<open>Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\<close>
+
+proposition contour_integral_uniform_limit:
+  assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on \<gamma>) F"
+      and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image \<gamma>. norm(f n x - l x) < e) F"
+      and noleB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
+      and \<gamma>: "valid_path \<gamma>"
+      and [simp]: "~ (trivial_limit F)"
+  shows "l contour_integrable_on \<gamma>" "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
+proof -
+  have "0 \<le> B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one)
+  { fix e::real
+    assume "0 < e"
+    then have eB: "0 < e / (\<bar>B\<bar> + 1)" by simp
+    obtain a where fga: "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (f a (\<gamma> x) - l (\<gamma> x)) < e / (\<bar>B\<bar> + 1)"
+               and inta: "(\<lambda>t. f a (\<gamma> t) * vector_derivative \<gamma> (at t)) integrable_on {0..1}"
+      using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]]
+      by (fastforce simp: contour_integrable_on path_image_def)
+    have Ble: "B * e / (\<bar>B\<bar> + 1) \<le> e"
+      using \<open>0 \<le> B\<close>  \<open>0 < e\<close> by (simp add: divide_simps)
+    have "\<exists>h. (\<forall>x\<in>{0..1}. cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - h x) \<le> e) \<and> h integrable_on {0..1}"
+      apply (rule_tac x="\<lambda>x. f (a::'a) (\<gamma> x) * vector_derivative \<gamma> (at x)" in exI)
+      apply (intro inta conjI ballI)
+      apply (rule order_trans [OF _ Ble])
+      apply (frule noleB)
+      apply (frule fga)
+      using \<open>0 \<le> B\<close>  \<open>0 < e\<close>
+      apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps)
+      apply (drule (1) mult_mono [OF less_imp_le])
+      apply (simp_all add: mult_ac)
+      done
+  }
+  then show lintg: "l contour_integrable_on \<gamma>"
+    apply (simp add: contour_integrable_on)
+    apply (blast intro: integrable_uniform_limit_real)
+    done
+  { fix e::real
+    define B' where "B' = B + 1"
+    have B': "B' > 0" "B' > B" using  \<open>0 \<le> B\<close> by (auto simp: B'_def)
+    assume "0 < e"
+    then have ev_no': "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image \<gamma>. 2 * cmod (f n x - l x) < e / B'"
+      using ev_no [of "e / B' / 2"] B' by (simp add: field_simps)
+    have ie: "integral {0..1::real} (\<lambda>x. e / 2) < e" using \<open>0 < e\<close> by simp
+    have *: "cmod (f x (\<gamma> t) * vector_derivative \<gamma> (at t) - l (\<gamma> t) * vector_derivative \<gamma> (at t)) \<le> e / 2"
+             if t: "t\<in>{0..1}" and leB': "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) < e / B'" for x t
+    proof -
+      have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) \<le> e * (B/ B')"
+        using mult_mono [OF less_imp_le [OF leB'] noleB] B' \<open>0 < e\<close> t by auto
+      also have "... < e"
+        by (simp add: B' \<open>0 < e\<close> mult_imp_div_pos_less)
+      finally have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) < e" .
+      then show ?thesis
+        by (simp add: left_diff_distrib [symmetric] norm_mult)
+    qed
+    have "\<forall>\<^sub>F x in F. dist (contour_integral \<gamma> (f x)) (contour_integral \<gamma> l) < e"
+      apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]])
+      apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral)
+      apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify)
+      apply (rule le_less_trans [OF integral_norm_bound_integral ie])
+      apply (simp add: lintg integrable_diff contour_integrable_on [symmetric])
+      apply (blast intro: *)+
+      done
+  }
+  then show "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
+    by (rule tendstoI)
+qed
+
+proposition contour_integral_uniform_limit_circlepath:
+  assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on (circlepath z r)) F"
+      and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image (circlepath z r). norm(f n x - l x) < e) F"
+      and [simp]: "~ (trivial_limit F)" "0 < r"
+  shows "l contour_integrable_on (circlepath z r)" "((\<lambda>n. contour_integral (circlepath z r) (f n)) \<longlongrightarrow> contour_integral (circlepath z r) l) F"
+by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms)
+
+
+subsection\<open> General stepping result for derivative formulas.\<close>
+
+lemma sum_sqs_eq:
+  fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \<Longrightarrow> y = x"
+  by algebra
+
+proposition Cauchy_next_derivative:
+  assumes "continuous_on (path_image \<gamma>) f'"
+      and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
+      and int: "\<And>w. w \<in> s - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>"
+      and k: "k \<noteq> 0"
+      and "open s"
+      and \<gamma>: "valid_path \<gamma>"
+      and w: "w \<in> s - path_image \<gamma>"
+    shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>"
+      and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k))))
+           (at w)"  (is "?thes2")
+proof -
+  have "open (s - path_image \<gamma>)" using \<open>open s\<close> closed_valid_path_image \<gamma> by blast
+  then obtain d where "d>0" and d: "ball w d \<subseteq> s - path_image \<gamma>" using w
+    using open_contains_ball by blast
+  have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n"
+    by (metis norm_of_nat of_nat_Suc)
+  have 1: "\<forall>\<^sub>F n in at w. (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k)
+                         contour_integrable_on \<gamma>"
+    apply (simp add: eventually_at)
+    apply (rule_tac x=d in exI)
+    apply (simp add: \<open>d > 0\<close> dist_norm field_simps, clarify)
+    apply (rule contour_integrable_div [OF contour_integrable_diff])
+    using int w d
+    apply (force simp:  dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+
+    done
+  have bim_g: "bounded (image f' (path_image \<gamma>))"
+    by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms)
+  then obtain C where "C > 0" and C: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cmod (f' (\<gamma> x)) \<le> C"
+    by (force simp: bounded_pos path_image_def)
+  have twom: "\<forall>\<^sub>F n in at w.
+               \<forall>x\<in>path_image \<gamma>.
+                cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e"
+         if "0 < e" for e
+  proof -
+    have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k)   < e"
+            if x: "x \<in> path_image \<gamma>" and "u \<noteq> w" and uwd: "cmod (u - w) < d/2"
+                and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)"
+            for u x
+    proof -
+      define ff where [abs_def]:
+        "ff n w =
+          (if n = 0 then inverse(x - w)^k
+           else if n = 1 then k / (x - w)^(Suc k)
+           else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w
+      have km1: "\<And>z::complex. z \<noteq> 0 \<Longrightarrow> z ^ (k - Suc 0) = z ^ k / z"
+        by (simp add: field_simps) (metis Suc_pred \<open>k \<noteq> 0\<close> neq0_conv power_Suc)
+      have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))"
+              if "z \<in> ball w (d / 2)" "i \<le> 1" for i z
+      proof -
+        have "z \<notin> path_image \<gamma>"
+          using \<open>x \<in> path_image \<gamma>\<close> d that ball_divide_subset_numeral by blast
+        then have xz[simp]: "x \<noteq> z" using \<open>x \<in> path_image \<gamma>\<close> by blast
+        then have neq: "x * x + z * z \<noteq> x * (z * 2)"
+          by (blast intro: dest!: sum_sqs_eq)
+        with xz have "\<And>v. v \<noteq> 0 \<Longrightarrow> (x * x + z * z) * v \<noteq> (x * (z * 2) * v)" by auto
+        then have neqq: "\<And>v. v \<noteq> 0 \<Longrightarrow> x * (x * v) + z * (z * v) \<noteq> x * (z * (2 * v))"
+          by (simp add: algebra_simps)
+        show ?thesis using \<open>i \<le> 1\<close>
+          apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe)
+          apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+
+          done
+      qed
+      { fix a::real and b::real assume ab: "a > 0" "b > 0"
+        then have "k * (1 + real k) * (1 / a) \<le> k * (1 + real k) * (4 / b) \<longleftrightarrow> b \<le> 4 * a"
+          apply (subst mult_le_cancel_left_pos)
+          using \<open>k \<noteq> 0\<close>
+          apply (auto simp: divide_simps)
+          done
+        with ab have "real k * (1 + real k) / a \<le> (real k * 4 + real k * real k * 4) / b \<longleftrightarrow> b \<le> 4 * a"
+          by (simp add: field_simps)
+      } note canc = this
+      have ff2: "cmod (ff (Suc 1) v) \<le> real (k * (k + 1)) / (d / 2) ^ (k + 2)"
+                if "v \<in> ball w (d / 2)" for v
+      proof -
+        have "d/2 \<le> cmod (x - v)" using d x that
+          apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify)
+          apply (drule subsetD)
+           prefer 2 apply blast
+          apply (metis norm_minus_commute norm_triangle_half_r CollectI)
+          done
+        then have "d \<le> cmod (x - v) * 2"
+          by (simp add: divide_simps)
+        then have dpow_le: "d ^ (k+2) \<le> (cmod (x - v) * 2) ^ (k+2)"
+          using \<open>0 < d\<close> order_less_imp_le power_mono by blast
+        have "x \<noteq> v" using that
+          using \<open>x \<in> path_image \<gamma>\<close> ball_divide_subset_numeral d by fastforce
+        then show ?thesis
+        using \<open>d > 0\<close>
+        apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc)
+        using dpow_le
+        apply (simp add: algebra_simps divide_simps mult_less_0_iff)
+        done
+      qed
+      have ub: "u \<in> ball w (d / 2)"
+        using uwd by (simp add: dist_commute dist_norm)
+      have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
+                  \<le> (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))"
+        using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified]
+        by (simp add: ff_def \<open>0 < d\<close>)
+      then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
+                  \<le> (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
+        by (simp add: field_simps)
+      then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
+                 / (cmod (u - w) * real k)
+                  \<le> (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
+        using \<open>k \<noteq> 0\<close> \<open>u \<noteq> w\<close> by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq)
+      also have "... < e"
+        using uw_less \<open>0 < d\<close> by (simp add: mult_ac divide_simps)
+      finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k)))
+                        / cmod ((u - w) * real k)   <   e"
+        by (simp add: norm_mult)
+      have "x \<noteq> u"
+        using uwd \<open>0 < d\<close> x d by (force simp: dist_norm ball_def norm_minus_commute)
+      show ?thesis
+        apply (rule le_less_trans [OF _ e])
+        using \<open>k \<noteq> 0\<close> \<open>x \<noteq> u\<close>  \<open>u \<noteq> w\<close>
+        apply (simp add: field_simps norm_divide [symmetric])
+        done
+    qed
+    show ?thesis
+      unfolding eventually_at
+      apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI)
+      apply (force simp: \<open>d > 0\<close> dist_norm that simp del: power_Suc intro: *)
+      done
+  qed
+  have 2: "\<forall>\<^sub>F n in at w.
+              \<forall>x\<in>path_image \<gamma>.
+               cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e"
+          if "0 < e" for e
+  proof -
+    have *: "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
+                        f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) < e"
+              if ec: "cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
+                      inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k) < e / C"
+                 and x: "0 \<le> x" "x \<le> 1"
+              for u x
+    proof (cases "(f' (\<gamma> x)) = 0")
+      case True then show ?thesis by (simp add: \<open>0 < e\<close>)
+    next
+      case False
+      have "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
+                        f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) =
+            cmod (f' (\<gamma> x) * ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
+                             inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k))"
+        by (simp add: field_simps)
+      also have "... = cmod (f' (\<gamma> x)) *
+                       cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
+                             inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k)"
+        by (simp add: norm_mult)
+      also have "... < cmod (f' (\<gamma> x)) * (e/C)"
+        apply (rule mult_strict_left_mono [OF ec])
+        using False by simp
+      also have "... \<le> e" using C
+        by (metis False \<open>0 < e\<close> frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff)
+      finally show ?thesis .
+    qed
+    show ?thesis
+      using twom [OF divide_pos_pos [OF that \<open>C > 0\<close>]]   unfolding path_image_def
+      by (force intro: * elim: eventually_mono)
+  qed
+  show "(\<lambda>u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \<gamma>"
+    by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
+  have *: "(\<lambda>n. contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k))
+           \<midarrow>w\<rightarrow> contour_integral \<gamma> (\<lambda>u. f' u / (u - w) ^ (Suc k))"
+    by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
+  have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) =
+              (f u - f w) / (u - w) / k"
+           if "dist u w < d" for u
+    apply (rule contour_integral_unique)
+    apply (simp add: diff_divide_distrib algebra_simps)
+    apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int)
+    apply (metis contra_subsetD d dist_commute mem_ball that)
+    apply (rule w)
+    done
+  show ?thes2
+    apply (simp add: DERIV_within_iff del: power_Suc)
+    apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ])
+    apply (simp add: \<open>k \<noteq> 0\<close> **)
+    done
+qed
+
+corollary Cauchy_next_derivative_circlepath:
+  assumes contf: "continuous_on (path_image (circlepath z r)) f"
+      and int: "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)"
+      and k: "k \<noteq> 0"
+      and w: "w \<in> ball z r"
+    shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
+           (is "?thes1")
+      and "(g has_field_derivative (k * contour_integral (circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)))) (at w)"
+           (is "?thes2")
+proof -
+  have "r > 0" using w
+    using ball_eq_empty by fastforce
+  have wim: "w \<in> ball z r - path_image (circlepath z r)"
+    using w by (auto simp: dist_norm)
+  show ?thes1 ?thes2
+    by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \<bar>r\<bar>"];
+        auto simp: vector_derivative_circlepath norm_mult)+
+qed
+
+
+text\<open> In particular, the first derivative formula.\<close>
+
+proposition Cauchy_derivative_integral_circlepath:
+  assumes contf: "continuous_on (cball z r) f"
+      and holf: "f holomorphic_on ball z r"
+      and w: "w \<in> ball z r"
+    shows "(\<lambda>u. f u/(u - w)^2) contour_integrable_on (circlepath z r)"
+           (is "?thes1")
+      and "(f has_field_derivative (1 / (2 * of_real pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u / (u - w)^2))) (at w)"
+           (is "?thes2")
+proof -
+  have [simp]: "r \<ge> 0" using w
+    using ball_eq_empty by fastforce
+  have f: "continuous_on (path_image (circlepath z r)) f"
+    by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def)
+  have int: "\<And>w. dist z w < r \<Longrightarrow>
+                 ((\<lambda>u. f u / (u - w)) has_contour_integral (\<lambda>x. 2 * of_real pi * \<i> * f x) w) (circlepath z r)"
+    by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute)
+  show ?thes1
+    apply (simp add: power2_eq_square)
+    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified])
+    apply (blast intro: int)
+    done
+  have "((\<lambda>x. 2 * of_real pi * \<i> * f x) has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2)) (at w)"
+    apply (simp add: power2_eq_square)
+    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x", simplified])
+    apply (blast intro: int)
+    done
+  then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2) / (2 * of_real pi * \<i>)) (at w)"
+    by (rule DERIV_cdivide [where f = "\<lambda>x. 2 * of_real pi * \<i> * f x" and c = "2 * of_real pi * \<i>", simplified])
+  show ?thes2
+    by simp (rule fder)
+qed
+
+subsection\<open> Existence of all higher derivatives.\<close>
+
+proposition derivative_is_holomorphic:
+  assumes "open s"
+      and fder: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z)"
+    shows "f' holomorphic_on s"
+proof -
+  have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> s" for z
+  proof -
+    obtain r where "r > 0" and r: "cball z r \<subseteq> s"
+      using open_contains_cball \<open>z \<in> s\<close> \<open>open s\<close> by blast
+    then have holf_cball: "f holomorphic_on cball z r"
+      apply (simp add: holomorphic_on_def)
+      using field_differentiable_at_within field_differentiable_def fder by blast
+    then have "continuous_on (path_image (circlepath z r)) f"
+      using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on])
+    then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)"
+      by (auto intro: continuous_intros)+
+    have contf_cball: "continuous_on (cball z r) f" using holf_cball
+      by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset)
+    have holf_ball: "f holomorphic_on ball z r" using holf_cball
+      using ball_subset_cball holomorphic_on_subset by blast
+    { fix w  assume w: "w \<in> ball z r"
+      have intf: "(\<lambda>u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r"
+        by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
+      have fder': "(f has_field_derivative 1 / (2 * of_real pi * \<i>) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2))
+                  (at w)"
+        by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
+      have f'_eq: "f' w = contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)"
+        using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder])
+      have "((\<lambda>u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \<i>)) has_contour_integral
+                contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
+                (circlepath z r)"
+        by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]])
+      then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral
+                contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
+                (circlepath z r)"
+        by (simp add: algebra_simps)
+      then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)"
+        by (simp add: f'_eq)
+    } note * = this
+    show ?thesis
+      apply (rule exI)
+      apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified])
+      apply (simp_all add: \<open>0 < r\<close> * dist_norm)
+      done
+  qed
+  show ?thesis
+    by (simp add: holomorphic_on_open [OF \<open>open s\<close>] *)
+qed
+
+lemma holomorphic_deriv [holomorphic_intros]:
+    "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv f) holomorphic_on s"
+by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def)
+
+lemma analytic_deriv: "f analytic_on s \<Longrightarrow> (deriv f) analytic_on s"
+  using analytic_on_holomorphic holomorphic_deriv by auto
+
+lemma holomorphic_higher_deriv [holomorphic_intros]: "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv ^^ n) f holomorphic_on s"
+  by (induction n) (auto simp: holomorphic_deriv)
+
+lemma analytic_higher_deriv: "f analytic_on s \<Longrightarrow> (deriv ^^ n) f analytic_on s"
+  unfolding analytic_on_def using holomorphic_higher_deriv by blast
+
+lemma has_field_derivative_higher_deriv:
+     "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk>
+      \<Longrightarrow> ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)"
+by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply
+         funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def)
+
+lemma valid_path_compose_holomorphic:
+  assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \<subseteq> s"
+  shows "valid_path (f o g)"
+proof (rule valid_path_compose[OF \<open>valid_path g\<close>])
+  fix x assume "x \<in> path_image g"
+  then show "\<exists>f'. (f has_field_derivative f') (at x)"
+    using holo holomorphic_on_open[OF \<open>open s\<close>] \<open>path_image g \<subseteq> s\<close> by auto
+next
+  have "deriv f holomorphic_on s"
+    using holomorphic_deriv holo \<open>open s\<close> by auto
+  then show "continuous_on (path_image g) (deriv f)"
+    using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto
+qed
+
+
+subsection\<open> Morera's theorem.\<close>
+
+lemma Morera_local_triangle_ball:
+  assumes "\<And>z. z \<in> s
+          \<Longrightarrow> \<exists>e a. 0 < e \<and> z \<in> ball a e \<and> continuous_on (ball a e) f \<and>
+                    (\<forall>b c. closed_segment b c \<subseteq> ball a e
+                           \<longrightarrow> contour_integral (linepath a b) f +
+                               contour_integral (linepath b c) f +
+                               contour_integral (linepath c a) f = 0)"
+  shows "f analytic_on s"
+proof -
+  { fix z  assume "z \<in> s"
+    with assms obtain e a where
+            "0 < e" and z: "z \<in> ball a e" and contf: "continuous_on (ball a e) f"
+        and 0: "\<And>b c. closed_segment b c \<subseteq> ball a e
+                      \<Longrightarrow> contour_integral (linepath a b) f +
+                          contour_integral (linepath b c) f +
+                          contour_integral (linepath c a) f = 0"
+      by fastforce
+    have az: "dist a z < e" using mem_ball z by blast
+    have sb_ball: "ball z (e - dist a z) \<subseteq> ball a e"
+      by (simp add: dist_commute ball_subset_ball_iff)
+    have "\<exists>e>0. f holomorphic_on ball z e"
+      apply (rule_tac x="e - dist a z" in exI)
+      apply (simp add: az)
+      apply (rule holomorphic_on_subset [OF _ sb_ball])
+      apply (rule derivative_is_holomorphic[OF open_ball])
+      apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a])
+         apply (simp_all add: 0 \<open>0 < e\<close>)
+      apply (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball)
+      done
+  }
+  then show ?thesis
+    by (simp add: analytic_on_def)
+qed
+
+lemma Morera_local_triangle:
+  assumes "\<And>z. z \<in> s
+          \<Longrightarrow> \<exists>t. open t \<and> z \<in> t \<and> continuous_on t f \<and>
+                  (\<forall>a b c. convex hull {a,b,c} \<subseteq> t
+                              \<longrightarrow> contour_integral (linepath a b) f +
+                                  contour_integral (linepath b c) f +
+                                  contour_integral (linepath c a) f = 0)"
+  shows "f analytic_on s"
+proof -
+  { fix z  assume "z \<in> s"
+    with assms obtain t where
+            "open t" and z: "z \<in> t" and contf: "continuous_on t f"
+        and 0: "\<And>a b c. convex hull {a,b,c} \<subseteq> t
+                      \<Longrightarrow> contour_integral (linepath a b) f +
+                          contour_integral (linepath b c) f +
+                          contour_integral (linepath c a) f = 0"
+      by force
+    then obtain e where "e>0" and e: "ball z e \<subseteq> t"
+      using open_contains_ball by blast
+    have [simp]: "continuous_on (ball z e) f" using contf
+      using continuous_on_subset e by blast
+    have "\<exists>e a. 0 < e \<and>
+               z \<in> ball a e \<and>
+               continuous_on (ball a e) f \<and>
+               (\<forall>b c. closed_segment b c \<subseteq> ball a e \<longrightarrow>
+                      contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)"
+      apply (rule_tac x=e in exI)
+      apply (rule_tac x=z in exI)
+      apply (simp add: \<open>e > 0\<close>, clarify)
+      apply (rule 0)
+      apply (meson z \<open>0 < e\<close> centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset)
+      done
+  }
+  then show ?thesis
+    by (simp add: Morera_local_triangle_ball)
+qed
+
+proposition Morera_triangle:
+    "\<lbrakk>continuous_on s f; open s;
+      \<And>a b c. convex hull {a,b,c} \<subseteq> s
+              \<longrightarrow> contour_integral (linepath a b) f +
+                  contour_integral (linepath b c) f +
+                  contour_integral (linepath c a) f = 0\<rbrakk>
+     \<Longrightarrow> f analytic_on s"
+  using Morera_local_triangle by blast
+
+
+
+subsection\<open> Combining theorems for higher derivatives including Leibniz rule.\<close>
+
+lemma higher_deriv_linear [simp]:
+    "(deriv ^^ n) (\<lambda>w. c*w) = (\<lambda>z. if n = 0 then c*z else if n = 1 then c else 0)"
+  by (induction n) (auto simp: deriv_const deriv_linear)
+
+lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)"
+  by (induction n) (auto simp: deriv_const)
+
+lemma higher_deriv_ident [simp]:
+     "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)"
+  apply (induction n, simp)
+  apply (metis higher_deriv_linear lambda_one)
+  done
+
+corollary higher_deriv_id [simp]:
+     "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)"
+  by (simp add: id_def)
+
+lemma has_complex_derivative_funpow_1:
+     "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)"
+  apply (induction n)
+  apply auto
+  apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one)
+  by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral)
+
+proposition higher_deriv_uminus:
+  assumes "f holomorphic_on s" "open s" and z: "z \<in> s"
+    shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
+using z
+proof (induction n arbitrary: z)
+  case 0 then show ?case by simp
+next
+  case (Suc n z)
+  have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
+    using Suc.prems assms has_field_derivative_higher_deriv by auto
+  show ?case
+    apply simp
+    apply (rule DERIV_imp_deriv)
+    apply (rule DERIV_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"])
+    apply (rule derivative_eq_intros | rule * refl assms Suc)+
+    apply (simp add: Suc)
+    done
+qed
+
+proposition higher_deriv_add:
+  fixes z::complex
+  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
+    shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
+using z
+proof (induction n arbitrary: z)
+  case 0 then show ?case by simp
+next
+  case (Suc n z)
+  have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
+          "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
+    using Suc.prems assms has_field_derivative_higher_deriv by auto
+  show ?case
+    apply simp
+    apply (rule DERIV_imp_deriv)
+    apply (rule DERIV_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"])
+    apply (rule derivative_eq_intros | rule * refl assms Suc)+
+    apply (simp add: Suc)
+    done
+qed
+
+corollary higher_deriv_diff:
+  fixes z::complex
+  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
+    shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
+  apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add)
+  apply (subst higher_deriv_add)
+  using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus)
+  done
+
+
+lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))"
+  by (cases k) simp_all
+
+proposition higher_deriv_mult:
+  fixes z::complex
+  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
+    shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
+           (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
+using z
+proof (induction n arbitrary: z)
+  case 0 then show ?case by simp
+next
+  case (Suc n z)
+  have *: "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
+          "\<And>n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
+    using Suc.prems assms has_field_derivative_higher_deriv by auto
+  have sumeq: "(\<Sum>i = 0..n.
+               of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) =
+            g z * deriv ((deriv ^^ n) f) z + (\<Sum>i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))"
+    apply (simp add: bb algebra_simps setsum.distrib)
+    apply (subst (4) setsum_Suc_reindex)
+    apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong)
+    done
+  show ?case
+    apply (simp only: funpow.simps o_apply)
+    apply (rule DERIV_imp_deriv)
+    apply (rule DERIV_transform_within_open
+             [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"])
+    apply (simp add: algebra_simps)
+    apply (rule DERIV_cong [OF DERIV_setsum])
+    apply (rule DERIV_cmult)
+    apply (auto simp: intro: DERIV_mult * sumeq \<open>open s\<close> Suc.prems Suc.IH [symmetric])
+    done
+qed
+
+
+proposition higher_deriv_transform_within_open:
+  fixes z::complex
+  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
+      and fg: "\<And>w. w \<in> s \<Longrightarrow> f w = g w"
+    shows "(deriv ^^ i) f z = (deriv ^^ i) g z"
+using z
+by (induction i arbitrary: z)
+   (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms)
+
+proposition higher_deriv_compose_linear:
+  fixes z::complex
+  assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \<in> s"
+      and fg: "\<And>w. w \<in> s \<Longrightarrow> u * w \<in> t"
+    shows "(deriv ^^ n) (\<lambda>w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)"
+using z
+proof (induction n arbitrary: z)
+  case 0 then show ?case by simp
+next
+  case (Suc n z)
+  have holo0: "f holomorphic_on op * u ` s"
+    by (meson fg f holomorphic_on_subset image_subset_iff)
+  have holo1: "(\<lambda>w. f (u * w)) holomorphic_on s"
+    apply (rule holomorphic_on_compose [where g=f, unfolded o_def])
+    apply (rule holo0 holomorphic_intros)+
+    done
+  have holo2: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s"
+    apply (rule holomorphic_intros)+
+    apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def])
+    apply (rule holomorphic_intros)
+    apply (rule holomorphic_on_subset [where s=t])
+    apply (rule holomorphic_intros assms)+
+    apply (blast intro: fg)
+    done
+  have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z"
+    apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems])
+    apply (rule holomorphic_higher_deriv [OF holo1 s])
+    apply (simp add: Suc.IH)
+    done
+  also have "... = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z"
+    apply (rule deriv_cmult)
+    apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems])
+    apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def])
+    apply (simp add: analytic_on_linear)
+    apply (simp add: analytic_on_open f holomorphic_higher_deriv t)
+    apply (blast intro: fg)
+    done
+  also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)"
+      apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def])
+      apply (rule derivative_intros)
+      using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast
+      apply (simp add: deriv_linear)
+      done
+  finally show ?case
+    by simp
+qed
+
+lemma higher_deriv_add_at:
+  assumes "f analytic_on {z}" "g analytic_on {z}"
+    shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
+proof -
+  have "f analytic_on {z} \<and> g analytic_on {z}"
+    using assms by blast
+  with higher_deriv_add show ?thesis
+    by (auto simp: analytic_at_two)
+qed
+
+lemma higher_deriv_diff_at:
+  assumes "f analytic_on {z}" "g analytic_on {z}"
+    shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
+proof -
+  have "f analytic_on {z} \<and> g analytic_on {z}"
+    using assms by blast
+  with higher_deriv_diff show ?thesis
+    by (auto simp: analytic_at_two)
+qed
+
+lemma higher_deriv_uminus_at:
+   "f analytic_on {z}  \<Longrightarrow> (deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
+  using higher_deriv_uminus
+    by (auto simp: analytic_at)
+
+lemma higher_deriv_mult_at:
+  assumes "f analytic_on {z}" "g analytic_on {z}"
+    shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
+           (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
+proof -
+  have "f analytic_on {z} \<and> g analytic_on {z}"
+    using assms by blast
+  with higher_deriv_mult show ?thesis
+    by (auto simp: analytic_at_two)
+qed
+
+
+text\<open> Nonexistence of isolated singularities and a stronger integral formula.\<close>
+
+proposition no_isolated_singularity:
+  fixes z::complex
+  assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k"
+    shows "f holomorphic_on s"
+proof -
+  { fix z
+    assume "z \<in> s" and cdf: "\<And>x. x\<in>s - k \<Longrightarrow> f field_differentiable at x"
+    have "f field_differentiable at z"
+    proof (cases "z \<in> k")
+      case False then show ?thesis by (blast intro: cdf \<open>z \<in> s\<close>)
+    next
+      case True
+      with finite_set_avoid [OF k, of z]
+      obtain d where "d>0" and d: "\<And>x. \<lbrakk>x\<in>k; x \<noteq> z\<rbrakk> \<Longrightarrow> d \<le> dist z x"
+        by blast
+      obtain e where "e>0" and e: "ball z e \<subseteq> s"
+        using  s \<open>z \<in> s\<close> by (force simp add: open_contains_ball)
+      have fde: "continuous_on (ball z (min d e)) f"
+        by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI)
+      have "\<exists>g. \<forall>w \<in> ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))"
+        apply (rule contour_integral_convex_primitive [OF convex_ball fde])
+        apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
+         apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset)
+        apply (rule cdf)
+        apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset
+               interior_mono interior_subset subset_hull)
+        done
+      then have "f holomorphic_on ball z (min d e)"
+        by (metis open_ball at_within_open derivative_is_holomorphic)
+      then show ?thesis
+        unfolding holomorphic_on_def
+        by (metis open_ball \<open>0 < d\<close> \<open>0 < e\<close> at_within_open centre_in_ball min_less_iff_conj)
+    qed
+  }
+  with holf s k show ?thesis
+    by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric])
+qed
+
+proposition Cauchy_integral_formula_convex:
+    assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f"
+        and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
+        and z: "z \<in> interior s" and vpg: "valid_path \<gamma>"
+        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
+  apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf])
+  apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def)
+  using no_isolated_singularity [where s = "interior s"]
+  apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset
+               has_field_derivative_at_within holomorphic_on_def interior_subset open_interior)
+  using assms
+  apply auto
+  done
+
+
+text\<open> Formula for higher derivatives.\<close>
+
+proposition Cauchy_has_contour_integral_higher_derivative_circlepath:
+  assumes contf: "continuous_on (cball z r) f"
+      and holf: "f holomorphic_on ball z r"
+      and w: "w \<in> ball z r"
+    shows "((\<lambda>u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \<i>) / (fact k) * (deriv ^^ k) f w))
+           (circlepath z r)"
+using w
+proof (induction k arbitrary: w)
+  case 0 then show ?case
+    using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm)
+next
+  case (Suc k)
+  have [simp]: "r > 0" using w
+    using ball_eq_empty by fastforce
+  have f: "continuous_on (path_image (circlepath z r)) f"
+    by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le)
+  obtain X where X: "((\<lambda>u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)"
+    using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems]
+    by (auto simp: contour_integrable_on_def)
+  then have con: "contour_integral (circlepath z r) ((\<lambda>u. f u / (u - w) ^ Suc (Suc k))) = X"
+    by (rule contour_integral_unique)
+  have "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)"
+    using Suc.prems assms has_field_derivative_higher_deriv by auto
+  then have dnf_diff: "\<And>n. (deriv ^^ n) f field_differentiable (at w)"
+    by (force simp add: field_differentiable_def)
+  have "deriv (\<lambda>w. complex_of_real (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w) w =
+          of_nat (Suc k) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w) ^ Suc (Suc k))"
+    by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems])
+  also have "... = of_nat (Suc k) * X"
+    by (simp only: con)
+  finally have "deriv (\<lambda>w. ((2 * pi) * \<i> / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" .
+  then have "((2 * pi) * \<i> / (fact k)) * deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X"
+    by (metis deriv_cmult dnf_diff)
+  then have "deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \<i> / (fact k))"
+    by (simp add: field_simps)
+  then show ?case
+  using of_nat_eq_0_iff X by fastforce
+qed
+
+proposition Cauchy_higher_derivative_integral_circlepath:
+  assumes contf: "continuous_on (cball z r) f"
+      and holf: "f holomorphic_on ball z r"
+      and w: "w \<in> ball z r"
+    shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
+           (is "?thes1")
+      and "(deriv ^^ k) f w = (fact k) / (2 * pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k))"
+           (is "?thes2")
+proof -
+  have *: "((\<lambda>u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w)
+           (circlepath z r)"
+    using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms]
+    by simp
+  show ?thes1 using *
+    using contour_integrable_on_def by blast
+  show ?thes2
+    unfolding contour_integral_unique [OF *] by (simp add: divide_simps)
+qed
+
+corollary Cauchy_contour_integral_circlepath:
+  assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
+    shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)) = (2 * pi * \<i>) * (deriv ^^ k) f w / (fact k)"
+by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms])
+
+corollary Cauchy_contour_integral_circlepath_2:
+  assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
+    shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^2) = (2 * pi * \<i>) * deriv f w"
+  using Cauchy_contour_integral_circlepath [OF assms, of 1]
+  by (simp add: power2_eq_square)
+
+
+subsection\<open>A holomorphic function is analytic, i.e. has local power series.\<close>
+
+theorem holomorphic_power_series:
+  assumes holf: "f holomorphic_on ball z r"
+      and w: "w \<in> ball z r"
+    shows "((\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
+proof -
+  have fh': "f holomorphic_on cball z ((r + dist w z) / 2)"
+     apply (rule holomorphic_on_subset [OF holf])
+     apply (clarsimp simp del: divide_const_simps)
+     apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w)
+     done
+  \<comment>\<open>Replacing @{term r} and the original (weak) premises\<close>
+  obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \<in> ball z r"
+    apply (rule that [of "(r + dist w z) / 2"])
+      apply (simp_all add: fh')
+     apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w)
+    apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w)
+    done
+  then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f"
+    using ball_subset_cball holomorphic_on_subset apply blast
+    by (simp add: holfc holomorphic_on_imp_continuous_on)
+  have cint: "\<And>k. (\<lambda>u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r"
+    apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
+    apply (simp add: \<open>0 < r\<close>)
+    done
+  obtain B where "0 < B" and B: "\<And>u. u \<in> cball z r \<Longrightarrow> norm(f u) \<le> B"
+    by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI)
+  obtain k where k: "0 < k" "k \<le> r" and wz_eq: "norm(w - z) = r - k"
+             and kle: "\<And>u. norm(u - z) = r \<Longrightarrow> k \<le> norm(u - w)"
+    apply (rule_tac k = "r - dist z w" in that)
+    using w
+    apply (auto simp: dist_norm norm_minus_commute)
+    by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
+  have *: "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>path_image (circlepath z r).
+                norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e"
+          if "0 < e" for e
+  proof -
+    have rr: "0 \<le> (r - k) / r" "(r - k) / r < 1" using  k by auto
+    obtain n where n: "((r - k) / r) ^ n < e / B * k"
+      using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force
+    have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e"
+         if "n \<le> N" and r: "r = dist z u"  for N u
+    proof -
+      have N: "((r - k) / r) ^ N < e / B * k"
+        apply (rule le_less_trans [OF power_decreasing n])
+        using  \<open>n \<le> N\<close> k by auto
+      have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)"
+        using \<open>0 < r\<close> r w by auto
+      have wzu_not1: "(w - z) / (u - z) \<noteq> 1"
+        by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
+      have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
+            = norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)"
+        unfolding setsum_left_distrib setsum_divide_distrib power_divide by (simp add: algebra_simps)
+      also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)"
+        using \<open>0 < B\<close>
+        apply (auto simp: geometric_sum [OF wzu_not1])
+        apply (simp add: field_simps norm_mult [symmetric])
+        done
+      also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)"
+        using \<open>0 < r\<close> r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute)
+      also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)"
+        by (simp add: algebra_simps)
+      also have "... = norm (w - z) ^ N * norm (f u) / r ^ N"
+        by (simp add: norm_mult norm_power norm_minus_commute)
+      also have "... \<le> (((r - k)/r)^N) * B"
+        using \<open>0 < r\<close> w k
+        apply (simp add: divide_simps)
+        apply (rule mult_mono [OF power_mono])
+        apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r)
+        done
+      also have "... < e * k"
+        using \<open>0 < B\<close> N by (simp add: divide_simps)
+      also have "... \<le> e * norm (u - w)"
+        using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm)
+      finally show ?thesis
+        by (simp add: divide_simps norm_divide del: power_Suc)
+    qed
+    with \<open>0 < r\<close> show ?thesis
+      by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def)
+  qed
+  have eq: "\<forall>\<^sub>F x in sequentially.
+             contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
+             (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
+    apply (rule eventuallyI)
+    apply (subst contour_integral_setsum, simp)
+    using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps)
+    apply (simp only: contour_integral_lmul cint algebra_simps)
+    done
+  have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
+        sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))"
+    unfolding sums_def
+    apply (rule Lim_transform_eventually [OF eq])
+    apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *])
+    apply (rule contour_integrable_setsum, simp)
+    apply (rule contour_integrable_lmul)
+    apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
+    using \<open>0 < r\<close>
+    apply auto
+    done
+  then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
+             sums (2 * of_real pi * \<i> * f w)"
+    using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]])
+  then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z)^k / (\<i> * (of_real pi * 2)))
+            sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))"
+    by (rule sums_divide)
+  then have "(\<lambda>n. (w - z) ^ n * contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc n) / (\<i> * (of_real pi * 2)))
+            sums f w"
+    by (simp add: field_simps)
+  then show ?thesis
+    by (simp add: field_simps \<open>0 < r\<close> Cauchy_higher_derivative_integral_circlepath [OF contf holf])
+qed
+
+
+subsection\<open>The Liouville theorem and the Fundamental Theorem of Algebra.\<close>
+
+text\<open> These weak Liouville versions don't even need the derivative formula.\<close>
+
+lemma Liouville_weak_0:
+  assumes holf: "f holomorphic_on UNIV" and inf: "(f \<longlongrightarrow> 0) at_infinity"
+    shows "f z = 0"
+proof (rule ccontr)
+  assume fz: "f z \<noteq> 0"
+  with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"]
+  obtain B where B: "\<And>x. B \<le> cmod x \<Longrightarrow> norm (f x) * 2 < cmod (f z)"
+    by (auto simp: dist_norm)
+  define R where "R = 1 + \<bar>B\<bar> + norm z"
+  have "R > 0" unfolding R_def
+  proof -
+    have "0 \<le> cmod z + \<bar>B\<bar>"
+      by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def)
+    then show "0 < 1 + \<bar>B\<bar> + cmod z"
+      by linarith
+  qed
+  have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)"
+    apply (rule Cauchy_integral_circlepath)
+    using \<open>R > 0\<close> apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+
+    done
+  have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 \<le> cmod (f z)" for x
+    apply (simp add: R_def)
+    apply (rule less_imp_le)
+    apply (rule B)
+    using norm_triangle_ineq4 [of x z]
+    apply (auto simp:)
+    done
+  with  \<open>R > 0\<close> fz show False
+    using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"]
+    by (auto simp: norm_mult norm_divide divide_simps)
+qed
+
+proposition Liouville_weak:
+  assumes "f holomorphic_on UNIV" and "(f \<longlongrightarrow> l) at_infinity"
+    shows "f z = l"
+  using Liouville_weak_0 [of "\<lambda>z. f z - l"]
+  by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero)
+
+
+proposition Liouville_weak_inverse:
+  assumes "f holomorphic_on UNIV" and unbounded: "\<And>B. eventually (\<lambda>x. norm (f x) \<ge> B) at_infinity"
+    obtains z where "f z = 0"
+proof -
+  { assume f: "\<And>z. f z \<noteq> 0"
+    have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV"
+      by (simp add: holomorphic_on_divide holomorphic_on_const assms f)
+    have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity"
+      apply (rule tendstoI [OF eventually_mono])
+      apply (rule_tac B="2/e" in unbounded)
+      apply (simp add: dist_norm norm_divide divide_simps mult_ac)
+      done
+    have False
+      using Liouville_weak_0 [OF 1 2] f by simp
+  }
+  then show ?thesis
+    using that by blast
+qed
+
+
+text\<open> In particular we get the Fundamental Theorem of Algebra.\<close>
+
+theorem fundamental_theorem_of_algebra:
+    fixes a :: "nat \<Rightarrow> complex"
+  assumes "a 0 = 0 \<or> (\<exists>i \<in> {1..n}. a i \<noteq> 0)"
+  obtains z where "(\<Sum>i\<le>n. a i * z^i) = 0"
+using assms
+proof (elim disjE bexE)
+  assume "a 0 = 0" then show ?thesis
+    by (auto simp: that [of 0])
+next
+  fix i
+  assume i: "i \<in> {1..n}" and nz: "a i \<noteq> 0"
+  have 1: "(\<lambda>z. \<Sum>i\<le>n. a i * z^i) holomorphic_on UNIV"
+    by (rule holomorphic_intros)+
+  show ?thesis
+    apply (rule Liouville_weak_inverse [OF 1])
+    apply (rule polyfun_extremal)
+    apply (rule nz)
+    using i that
+    apply (auto simp:)
+    done
+qed
+
+
+subsection\<open> Weierstrass convergence theorem.\<close>
+
+proposition holomorphic_uniform_limit:
+  assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and> (f n) holomorphic_on ball z r) F"
+      and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
+      and F:  "~ trivial_limit F"
+  obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r"
+proof (cases r "0::real" rule: linorder_cases)
+  case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that)
+next
+  case equal then show ?thesis
+    by (force simp add: holomorphic_on_def continuous_on_sing intro: that)
+next
+  case greater
+  have contg: "continuous_on (cball z r) g"
+    using cont
+    by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F])
+  have 1: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1 / (2 * complex_of_real pi * \<i>) * g x)"
+    apply (rule continuous_intros continuous_on_subset [OF contg])+
+    using \<open>0 < r\<close> by auto
+  have 2: "((\<lambda>u. 1 / (2 * of_real pi * \<i>) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)"
+       if w: "w \<in> ball z r" for w
+  proof -
+    define d where "d = (r - norm(w - z))"
+    have "0 < d"  "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
+    have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)"
+      unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
+    have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r"
+      apply (rule eventually_mono [OF cont])
+      using w
+      apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified])
+      done
+    have ev_less: "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e"
+         if "e > 0" for e
+      using greater \<open>0 < d\<close> \<open>0 < e\<close>
+      apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps)
+      apply (rule_tac e1="e * d" in eventually_mono [OF lim])
+      apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+
+      done
+    have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r"
+      by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
+    have cif_tends_cig: "((\<lambda>n. contour_integral(circlepath z r) (\<lambda>u. f n u / (u - w))) \<longlongrightarrow> contour_integral(circlepath z r) (\<lambda>u. g u/(u - w))) F"
+      by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
+    have f_tends_cig: "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> contour_integral (circlepath z r) (\<lambda>u. g u / (u - w))) F"
+      apply (rule Lim_transform_eventually [where f = "\<lambda>n. contour_integral (circlepath z r) (\<lambda>u. f n u/(u - w))"])
+      apply (rule eventually_mono [OF cont])
+      apply (rule contour_integral_unique [OF Cauchy_integral_circlepath])
+      using w
+      apply (auto simp: norm_minus_commute dist_norm cif_tends_cig)
+      done
+    have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F"
+      apply (rule tendsto_mult_left [OF tendstoI])
+      apply (rule eventually_mono [OF lim], assumption)
+      using w
+      apply (force simp add: dist_norm)
+      done
+    then have "((\<lambda>u. g u / (u - w)) has_contour_integral 2 * of_real pi * \<i> * g w) (circlepath z r)"
+      using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w
+      by (force simp add: dist_norm)
+    then have "((\<lambda>u. g u / (2 * of_real pi * \<i> * (u - w))) has_contour_integral g w) (circlepath z r)"
+      using has_contour_integral_div [where c = "2 * of_real pi * \<i>"]
+      by (force simp add: field_simps)
+    then show ?thesis
+      by (simp add: dist_norm)
+  qed
+  show ?thesis
+    using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified]
+    by (fastforce simp add: holomorphic_on_open contg intro: that)
+qed
+
+
+text\<open> Version showing that the limit is the limit of the derivatives.\<close>
+
+proposition has_complex_derivative_uniform_limit:
+  fixes z::complex
+  assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and>
+                               (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F"
+      and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
+      and F:  "~ trivial_limit F" and "0 < r"
+  obtains g' where
+      "continuous_on (cball z r) g"
+      "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
+proof -
+  let ?conint = "contour_integral (circlepath z r)"
+  have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r"
+    by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F];
+             auto simp: holomorphic_on_open field_differentiable_def)+
+  then obtain g' where g': "\<And>x. x \<in> ball z r \<Longrightarrow> (g has_field_derivative g' x) (at x)"
+    using DERIV_deriv_iff_has_field_derivative
+    by (fastforce simp add: holomorphic_on_open)
+  then have derg: "\<And>x. x \<in> ball z r \<Longrightarrow> deriv g x = g' x"
+    by (simp add: DERIV_imp_deriv)
+  have tends_f'n_g': "((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" if w: "w \<in> ball z r" for w
+  proof -
+    have eq_f': "?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \<i>)"
+             if cont_fn: "continuous_on (cball z r) (f n)"
+             and fnd: "\<And>w. w \<in> ball z r \<Longrightarrow> (f n has_field_derivative f' n w) (at w)" for n
+    proof -
+      have hol_fn: "f n holomorphic_on ball z r"
+        using fnd by (force simp add: holomorphic_on_open)
+      have "(f n has_field_derivative 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)) (at w)"
+        by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w])
+      then have f': "f' n w = 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)"
+        using DERIV_unique [OF fnd] w by blast
+      show ?thesis
+        by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps)
+    qed
+    define d where "d = (r - norm(w - z))^2"
+    have "d > 0"
+      using w by (simp add: dist_commute dist_norm d_def)
+    have dle: "\<And>y. r = cmod (z - y) \<Longrightarrow> d \<le> cmod ((y - w)\<^sup>2)"
+      apply (simp add: d_def norm_power)
+      apply (rule power_mono)
+      apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
+      apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w)
+      done
+    have 1: "\<forall>\<^sub>F n in F. (\<lambda>x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r"
+      by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont])
+    have 2: "0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. \<forall>x \<in> path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e
+      using \<open>r > 0\<close>
+      apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def)
+      apply (rule eventually_mono [OF lim, of "e*d"])
+      apply (simp add: \<open>0 < d\<close>)
+      apply (force simp add: dist_norm dle intro: less_le_trans)
+      done
+    have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2))
+             \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F"
+      by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>])
+    then have tendsto_0: "((\<lambda>n. 1 / (2 * of_real pi * \<i>) * (?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2))) \<longlongrightarrow> 0) F"
+      using Lim_null by (force intro!: tendsto_mult_right_zero)
+    have "((\<lambda>n. f' n w - g' w) \<longlongrightarrow> 0) F"
+      apply (rule Lim_transform_eventually [OF _ tendsto_0])
+      apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont])
+      done
+    then show ?thesis using Lim_null by blast
+  qed
+  obtain g' where "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
+      by (blast intro: tends_f'n_g' g' )
+  then show ?thesis using g
+    using that by blast
+qed
+
+
+subsection\<open>Some more simple/convenient versions for applications.\<close>
+
+lemma holomorphic_uniform_sequence:
+  assumes s: "open s"
+      and hol_fn: "\<And>n. (f n) holomorphic_on s"
+      and to_g: "\<And>x. x \<in> s
+                     \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
+                             (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
+  shows "g holomorphic_on s"
+proof -
+  have "\<exists>f'. (g has_field_derivative f') (at z)" if "z \<in> s" for z
+  proof -
+    obtain r where "0 < r" and r: "cball z r \<subseteq> s"
+               and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
+      using to_g [OF \<open>z \<in> s\<close>] by blast
+    have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and> f n holomorphic_on ball z r"
+      apply (intro eventuallyI conjI)
+      using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast
+      apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r)
+      done
+    show ?thesis
+      apply (rule holomorphic_uniform_limit [OF *])
+      using \<open>0 < r\<close> centre_in_ball fg
+      apply (auto simp: holomorphic_on_open)
+      done
+  qed
+  with s show ?thesis
+    by (simp add: holomorphic_on_open)
+qed
+
+lemma has_complex_derivative_uniform_sequence:
+  fixes s :: "complex set"
+  assumes s: "open s"
+      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_field_derivative f' n x) (at x)"
+      and to_g: "\<And>x. x \<in> s
+             \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
+                     (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
+  shows "\<exists>g'. \<forall>x \<in> s. (g has_field_derivative g' x) (at x) \<and> ((\<lambda>n. f' n x) \<longlongrightarrow> g' x) sequentially"
+proof -
+  have y: "\<exists>y. (g has_field_derivative y) (at z) \<and> (\<lambda>n. f' n z) \<longlonglongrightarrow> y" if "z \<in> s" for z
+  proof -
+    obtain r where "0 < r" and r: "cball z r \<subseteq> s"
+               and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
+      using to_g [OF \<open>z \<in> s\<close>] by blast
+    have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and>
+                                   (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))"
+      apply (intro eventuallyI conjI)
+      apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s)
+      using ball_subset_cball hfd r apply blast
+      done
+    show ?thesis
+      apply (rule has_complex_derivative_uniform_limit [OF *, of g])
+      using \<open>0 < r\<close> centre_in_ball fg
+      apply force+
+      done
+  qed
+  show ?thesis
+    by (rule bchoice) (blast intro: y)
+qed
+
+
+subsection\<open>On analytic functions defined by a series.\<close>
+
+lemma series_and_derivative_comparison:
+  fixes s :: "complex set"
+  assumes s: "open s"
+      and h: "summable h"
+      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+      and to_g: "\<And>n x. \<lbrakk>N \<le> n; x \<in> s\<rbrakk> \<Longrightarrow> norm(f n x) \<le> h n"
+  obtains g g' where "\<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
+proof -
+  obtain g where g: "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n x. N \<le> n \<and> x \<in> s \<longrightarrow> dist (\<Sum>n<n. f n x) (g x) < e"
+    using series_comparison_uniform [OF h to_g, of N s] by force
+  have *: "\<exists>d>0. cball x d \<subseteq> s \<and> (\<forall>e>0. \<forall>\<^sub>F n in sequentially. \<forall>y\<in>cball x d. cmod ((\<Sum>a<n. f a y) - g y) < e)"
+         if "x \<in> s" for x
+  proof -
+    obtain d where "d>0" and d: "cball x d \<subseteq> s"
+      using open_contains_cball [of "s"] \<open>x \<in> s\<close> s by blast
+    then show ?thesis
+      apply (rule_tac x=d in exI)
+      apply (auto simp: dist_norm eventually_sequentially)
+      apply (metis g contra_subsetD dist_norm)
+      done
+  qed
+  have "(\<forall>x\<in>s. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> g x)"
+    using g by (force simp add: lim_sequentially)
+  moreover have "\<exists>g'. \<forall>x\<in>s. (g has_field_derivative g' x) (at x) \<and> (\<lambda>n. \<Sum>i<n. f' i x) \<longlonglongrightarrow> g' x"
+    by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+
+  ultimately show ?thesis
+    by (force simp add: sums_def  conj_commute intro: that)
+qed
+
+text\<open>A version where we only have local uniform/comparative convergence.\<close>
+
+lemma series_and_derivative_comparison_local:
+  fixes s :: "complex set"
+  assumes s: "open s"
+      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+      and to_g: "\<And>x. x \<in> s \<Longrightarrow>
+                      \<exists>d h N. 0 < d \<and> summable h \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> norm(f n y) \<le> h n)"
+  shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
+proof -
+  have "\<exists>y. (\<lambda>n. f n z) sums (\<Sum>n. f n z) \<and> (\<lambda>n. f' n z) sums y \<and> ((\<lambda>x. \<Sum>n. f n x) has_field_derivative y) (at z)"
+       if "z \<in> s" for z
+  proof -
+    obtain d h N where "0 < d" "summable h" and le_h: "\<And>n y. \<lbrakk>N \<le> n; y \<in> ball z d\<rbrakk> \<Longrightarrow> norm(f n y) \<le> h n"
+      using to_g \<open>z \<in> s\<close> by meson
+    then obtain r where "r>0" and r: "ball z r \<subseteq> ball z d \<inter> s" using \<open>z \<in> s\<close> s
+      by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq)
+    have 1: "open (ball z d \<inter> s)"
+      by (simp add: open_Int s)
+    have 2: "\<And>n x. x \<in> ball z d \<inter> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+      by (auto simp: hfd)
+    obtain g g' where gg': "\<forall>x \<in> ball z d \<inter> s. ((\<lambda>n. f n x) sums g x) \<and>
+                                    ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
+      by (auto intro: le_h series_and_derivative_comparison [OF 1 \<open>summable h\<close> hfd])
+    then have "(\<lambda>n. f' n z) sums g' z"
+      by (meson \<open>0 < r\<close> centre_in_ball contra_subsetD r)
+    moreover have "(\<lambda>n. f n z) sums (\<Sum>n. f n z)"
+      by (metis summable_comparison_test' summable_sums centre_in_ball \<open>0 < d\<close> \<open>summable h\<close> le_h)
+    moreover have "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' z) (at z)"
+      apply (rule_tac f=g in DERIV_transform_at [OF _ \<open>0 < r\<close>])
+      apply (simp add: gg' \<open>z \<in> s\<close> \<open>0 < d\<close>)
+      apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique)
+      done
+    ultimately show ?thesis by auto
+  qed
+  then show ?thesis
+    by (rule_tac x="\<lambda>x. suminf  (\<lambda>n. f n x)" in exI) meson
+qed
+
+
+text\<open>Sometimes convenient to compare with a complex series of positive reals. (?)\<close>
+
+lemma series_and_derivative_comparison_complex:
+  fixes s :: "complex set"
+  assumes s: "open s"
+      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
+      and to_g: "\<And>x. x \<in> s \<Longrightarrow>
+                      \<exists>d h N. 0 < d \<and> summable h \<and> range h \<subseteq> nonneg_Reals \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> cmod(f n y) \<le> cmod (h n))"
+  shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
+apply (rule series_and_derivative_comparison_local [OF s hfd], assumption)
+apply (frule to_g)
+apply (erule ex_forward)
+apply (erule exE)
+apply (rule_tac x="Re o h" in exI)
+apply (erule ex_forward)
+apply (simp add: summable_Re o_def )
+apply (elim conjE all_forward)
+apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff)
+done
+
+
+text\<open>In particular, a power series is analytic inside circle of convergence.\<close>
+
+lemma power_series_and_derivative_0:
+  fixes a :: "nat \<Rightarrow> complex" and r::real
+  assumes "summable (\<lambda>n. a n * r^n)"
+    shows "\<exists>g g'. \<forall>z. cmod z < r \<longrightarrow>
+             ((\<lambda>n. a n * z^n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * z^(n - 1)) sums g' z) \<and> (g has_field_derivative g' z) (at z)"
+proof (cases "0 < r")
+  case True
+    have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)"
+      by (rule derivative_eq_intros | simp)+
+    have y_le: "\<lbrakk>cmod (z - y) * 2 < r - cmod z\<rbrakk> \<Longrightarrow> cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" for z y
+      using \<open>r > 0\<close>
+      apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add)
+      using norm_triangle_ineq2 [of y z]
+      apply (simp only: diff_le_eq norm_minus_commute mult_2)
+      done
+    have "summable (\<lambda>n. a n * complex_of_real r ^ n)"
+      using assms \<open>r > 0\<close> by simp
+    moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
+      using \<open>r > 0\<close>
+      by (simp add: of_real_add [symmetric] del: of_real_add)
+    ultimately have sum: "\<And>z. cmod z < r \<Longrightarrow> summable (\<lambda>n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)"
+      by (rule power_series_conv_imp_absconv_weak)
+    have "\<exists>g g'. \<forall>z \<in> ball 0 r. (\<lambda>n.  (a n) * z ^ n) sums g z \<and>
+               (\<lambda>n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \<and> (g has_field_derivative g' z) (at z)"
+      apply (rule series_and_derivative_comparison_complex [OF open_ball der])
+      apply (rule_tac x="(r - norm z)/2" in exI)
+      apply (simp add: dist_norm)
+      apply (rule_tac x="\<lambda>n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI)
+      using \<open>r > 0\<close>
+      apply (auto simp: sum nonneg_Reals_divide_I)
+      apply (rule_tac x=0 in exI)
+      apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le)
+      done
+  then show ?thesis
+    by (simp add: dist_0_norm ball_def)
+next
+  case False then show ?thesis
+    apply (simp add: not_less)
+    using less_le_trans norm_not_less_zero by blast
+qed
+
+proposition power_series_and_derivative:
+  fixes a :: "nat \<Rightarrow> complex" and r::real
+  assumes "summable (\<lambda>n. a n * r^n)"
+    obtains g g' where "\<forall>z \<in> ball w r.
+             ((\<lambda>n. a n * (z - w) ^ n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \<and>
+              (g has_field_derivative g' z) (at z)"
+  using power_series_and_derivative_0 [OF assms]
+  apply clarify
+  apply (rule_tac g="(\<lambda>z. g(z - w))" in that)
+  using DERIV_shift [where z="-w"]
+  apply (auto simp: norm_minus_commute Ball_def dist_norm)
+  done
+
+proposition power_series_holomorphic:
+  assumes "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>n. a n*(w - z)^n) sums f w)"
+    shows "f holomorphic_on ball z r"
+proof -
+  have "\<exists>f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w
+  proof -
+    have inb: "z + complex_of_real ((dist z w + r) / 2) \<in> ball z r"
+    proof -
+      have wz: "cmod (w - z) < r" using w
+        by (auto simp: divide_simps dist_norm norm_minus_commute)
+      then have "0 \<le> r"
+        by (meson less_eq_real_def norm_ge_zero order_trans)
+      show ?thesis
+        using w by (simp add: dist_norm \<open>0\<le>r\<close> of_real_add [symmetric] del: of_real_add)
+    qed
+    have sum: "summable (\<lambda>n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))"
+      using assms [OF inb] by (force simp add: summable_def dist_norm)
+    obtain g g' where gg': "\<And>u. u \<in> ball z ((cmod (z - w) + r) / 2) \<Longrightarrow>
+                               (\<lambda>n. a n * (u - z) ^ n) sums g u \<and>
+                               (\<lambda>n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \<and> (g has_field_derivative g' u) (at u)"
+      by (rule power_series_and_derivative [OF sum, of z]) fastforce
+    have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u
+    proof -
+      have less: "cmod (z - u) * 2 < cmod (z - w) + r"
+        using that dist_triangle2 [of z u w]
+        by (simp add: dist_norm [symmetric] algebra_simps)
+      show ?thesis
+        apply (rule sums_unique2 [of "\<lambda>n. a n*(u - z)^n"])
+        using gg' [of u] less w
+        apply (auto simp: assms dist_norm)
+        done
+    qed
+    show ?thesis
+      apply (rule_tac x="g' w" in exI)
+      apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"])
+      using w gg' [of w]
+      apply (auto simp: dist_norm)
+      done
+  qed
+  then show ?thesis by (simp add: holomorphic_on_open)
+qed
+
+corollary holomorphic_iff_power_series:
+     "f holomorphic_on ball z r \<longleftrightarrow>
+      (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
+  apply (intro iffI ballI)
+   using holomorphic_power_series  apply force
+  apply (rule power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"])
+  apply force
+  done
+
+corollary power_series_analytic:
+     "(\<And>w. w \<in> ball z r \<Longrightarrow> (\<lambda>n. a n*(w - z)^n) sums f w) \<Longrightarrow> f analytic_on ball z r"
+  by (force simp add: analytic_on_open intro!: power_series_holomorphic)
+
+corollary analytic_iff_power_series:
+     "f analytic_on ball z r \<longleftrightarrow>
+      (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
+  by (simp add: analytic_on_open holomorphic_iff_power_series)
+
+
+subsection\<open>Equality between holomorphic functions, on open ball then connected set.\<close>
+
+lemma holomorphic_fun_eq_on_ball:
+   "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
+     w \<in> ball z r;
+     \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk>
+     \<Longrightarrow> f w = g w"
+  apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
+  apply (auto simp: holomorphic_iff_power_series)
+  done
+
+lemma holomorphic_fun_eq_0_on_ball:
+   "\<lbrakk>f holomorphic_on ball z r;  w \<in> ball z r;
+     \<And>n. (deriv ^^ n) f z = 0\<rbrakk>
+     \<Longrightarrow> f w = 0"
+  apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
+  apply (auto simp: holomorphic_iff_power_series)
+  done
+
+lemma holomorphic_fun_eq_0_on_connected:
+  assumes holf: "f holomorphic_on s" and "open s"
+      and cons: "connected s"
+      and der: "\<And>n. (deriv ^^ n) f z = 0"
+      and "z \<in> s" "w \<in> s"
+    shows "f w = 0"
+proof -
+  have *: "\<And>x e. \<lbrakk> \<forall>xa. (deriv ^^ xa) f x = 0;  ball x e \<subseteq> s\<rbrakk>
+           \<Longrightarrow> ball x e \<subseteq> (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
+    apply auto
+    apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv])
+    apply (rule holomorphic_on_subset [OF holf], simp_all)
+    by (metis funpow_add o_apply)
+  have 1: "openin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
+    apply (rule open_subset, force)
+    using \<open>open s\<close>
+    apply (simp add: open_contains_ball Ball_def)
+    apply (erule all_forward)
+    using "*" by auto blast+
+  have 2: "closedin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
+    using assms
+    by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
+  obtain e where "e>0" and e: "ball w e \<subseteq> s" using openE [OF \<open>open s\<close> \<open>w \<in> s\<close>] .
+  then have holfb: "f holomorphic_on ball w e"
+    using holf holomorphic_on_subset by blast
+  have 3: "(\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}) = s \<Longrightarrow> f w = 0"
+    using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb])
+  show ?thesis
+    using cons der \<open>z \<in> s\<close>
+    apply (simp add: connected_clopen)
+    apply (drule_tac x="\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}" in spec)
+    apply (auto simp: 1 2 3)
+    done
+qed
+
+lemma holomorphic_fun_eq_on_connected:
+  assumes "f holomorphic_on s" "g holomorphic_on s" and "open s"  "connected s"
+      and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z"
+      and "z \<in> s" "w \<in> s"
+    shows "f w = g w"
+  apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>x. f x - g x" s z, simplified])
+  apply (intro assms holomorphic_intros)
+  using assms apply simp_all
+  apply (subst higher_deriv_diff, auto)
+  done
+
+lemma holomorphic_fun_eq_const_on_connected:
+  assumes holf: "f holomorphic_on s" and "open s"
+      and cons: "connected s"
+      and der: "\<And>n. 0 < n \<Longrightarrow> (deriv ^^ n) f z = 0"
+      and "z \<in> s" "w \<in> s"
+    shows "f w = f z"
+  apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>w. f w - f z" s z, simplified])
+  apply (intro assms holomorphic_intros)
+  using assms apply simp_all
+  apply (subst higher_deriv_diff)
+  apply (intro holomorphic_intros | simp)+
+  done
+
+
+subsection\<open>Some basic lemmas about poles/singularities.\<close>
+
+lemma pole_lemma:
+  assumes holf: "f holomorphic_on s" and a: "a \<in> interior s"
+    shows "(\<lambda>z. if z = a then deriv f a
+                 else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s")
+proof -
+  have F1: "?F field_differentiable (at u within s)" if "u \<in> s" "u \<noteq> a" for u
+  proof -
+    have fcd: "f field_differentiable at u within s"
+      using holf holomorphic_on_def by (simp add: \<open>u \<in> s\<close>)
+    have cd: "(\<lambda>z. (f z - f a) / (z - a)) field_differentiable at u within s"
+      by (rule fcd derivative_intros | simp add: that)+
+    have "0 < dist a u" using that dist_nz by blast
+    then show ?thesis
+      by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \<open>u \<in> s\<close>)
+  qed
+  have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> s" for e
+  proof -
+    have holfb: "f holomorphic_on ball a e"
+      by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> s\<close>])
+    have 2: "?F holomorphic_on ball a e - {a}"
+      apply (rule holomorphic_on_subset [where s = "s - {a}"])
+      apply (simp add: holomorphic_on_def field_differentiable_def [symmetric])
+      using mem_ball that
+      apply (auto intro: F1 field_differentiable_within_subset)
+      done
+    have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x"
+            if "dist a x < e" for x
+    proof (cases "x=a")
+      case True then show ?thesis
+      using holfb \<open>0 < e\<close>
+      apply (simp add: holomorphic_on_open field_differentiable_def [symmetric])
+      apply (drule_tac x=a in bspec)
+      apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2
+                elim: rev_iffD1 [OF _ LIM_equal])
+      done
+    next
+      case False with 2 that show ?thesis
+        by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at)
+    qed
+    then have 1: "continuous_on (ball a e) ?F"
+      by (clarsimp simp:  continuous_on_eq_continuous_at)
+    have "?F holomorphic_on ball a e"
+      by (auto intro: no_isolated_singularity [OF 1 2])
+    with that show ?thesis
+      by (simp add: holomorphic_on_open field_differentiable_def [symmetric]
+                    field_differentiable_at_within)
+  qed
+  show ?thesis
+  proof
+    fix x assume "x \<in> s" show "?F field_differentiable at x within s"
+    proof (cases "x=a")
+      case True then show ?thesis
+      using a by (auto simp: mem_interior intro: field_differentiable_at_within F2)
+    next
+      case False with F1 \<open>x \<in> s\<close>
+      show ?thesis by blast
+    qed
+  qed
+qed
+
+proposition pole_theorem:
+  assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
+      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
+    shows "(\<lambda>z. if z = a then deriv g a
+                 else f z - g a/(z - a)) holomorphic_on s"
+  using pole_lemma [OF holg a]
+  by (rule holomorphic_transform) (simp add: eq divide_simps)
+
+lemma pole_lemma_open:
+  assumes "f holomorphic_on s" "open s"
+    shows "(\<lambda>z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s"
+proof (cases "a \<in> s")
+  case True with assms interior_eq pole_lemma
+    show ?thesis by fastforce
+next
+  case False with assms show ?thesis
+    apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify)
+    apply (rule field_differentiable_transform_within [where f = "\<lambda>z. (f z - f a)/(z - a)" and d = 1])
+    apply (rule derivative_intros | force)+
+    done
+qed
+
+proposition pole_theorem_open:
+  assumes holg: "g holomorphic_on s" and s: "open s"
+      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
+    shows "(\<lambda>z. if z = a then deriv g a
+                 else f z - g a/(z - a)) holomorphic_on s"
+  using pole_lemma_open [OF holg s]
+  by (rule holomorphic_transform) (auto simp: eq divide_simps)
+
+proposition pole_theorem_0:
+  assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
+      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
+      and [simp]: "f a = deriv g a" "g a = 0"
+    shows "f holomorphic_on s"
+  using pole_theorem [OF holg a eq]
+  by (rule holomorphic_transform) (auto simp: eq divide_simps)
+
+proposition pole_theorem_open_0:
+  assumes holg: "g holomorphic_on s" and s: "open s"
+      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
+      and [simp]: "f a = deriv g a" "g a = 0"
+    shows "f holomorphic_on s"
+  using pole_theorem_open [OF holg s eq]
+  by (rule holomorphic_transform) (auto simp: eq divide_simps)
+
+lemma pole_theorem_analytic:
+  assumes g: "g analytic_on s"
+      and eq: "\<And>z. z \<in> s
+             \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
+    shows "(\<lambda>z. if z = a then deriv g a
+                 else f z - g a/(z - a)) analytic_on s"
+using g
+apply (simp add: analytic_on_def Ball_def)
+apply (safe elim!: all_forward dest!: eq)
+apply (rule_tac x="min d e" in exI, simp)
+apply (rule pole_theorem_open)
+apply (auto simp: holomorphic_on_subset subset_ball)
+done
+
+lemma pole_theorem_analytic_0:
+  assumes g: "g analytic_on s"
+      and eq: "\<And>z. z \<in> s \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
+      and [simp]: "f a = deriv g a" "g a = 0"
+    shows "f analytic_on s"
+proof -
+  have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
+    by auto
+  show ?thesis
+    using pole_theorem_analytic [OF g eq] by simp
+qed
+
+lemma pole_theorem_analytic_open_superset:
+  assumes g: "g analytic_on s" and "s \<subseteq> t" "open t"
+      and eq: "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
+    shows "(\<lambda>z. if z = a then deriv g a
+                 else f z - g a/(z - a)) analytic_on s"
+  apply (rule pole_theorem_analytic [OF g])
+  apply (rule openE [OF \<open>open t\<close>])
+  using assms eq by auto
+
+lemma pole_theorem_analytic_open_superset_0:
+  assumes g: "g analytic_on s" "s \<subseteq> t" "open t" "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
+      and [simp]: "f a = deriv g a" "g a = 0"
+    shows "f analytic_on s"
+proof -
+  have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
+    by auto
+  have "(\<lambda>z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s"
+    by (rule pole_theorem_analytic_open_superset [OF g])
+  then show ?thesis by simp
+qed
+
+
+
+subsection\<open>General, homology form of Cauchy's theorem.\<close>
+
+text\<open>Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\<close>
+
+lemma contour_integral_continuous_on_linepath_2D:
+  assumes "open u" and cont_dw: "\<And>w. w \<in> u \<Longrightarrow> F w contour_integrable_on (linepath a b)"
+      and cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). F x y)"
+      and abu: "closed_segment a b \<subseteq> u"
+    shows "continuous_on u (\<lambda>w. contour_integral (linepath a b) (F w))"
+proof -
+  have *: "\<exists>d>0. \<forall>x'\<in>u. dist x' w < d \<longrightarrow>
+                         dist (contour_integral (linepath a b) (F x'))
+                              (contour_integral (linepath a b) (F w)) \<le> \<epsilon>"
+          if "w \<in> u" "0 < \<epsilon>" "a \<noteq> b" for w \<epsilon>
+  proof -
+    obtain \<delta> where "\<delta>>0" and \<delta>: "cball w \<delta> \<subseteq> u" using open_contains_cball \<open>open u\<close> \<open>w \<in> u\<close> by force
+    let ?TZ = "{(t,z) |t z. t \<in> cball w \<delta> \<and> z \<in> closed_segment a b}"
+    have "uniformly_continuous_on ?TZ (\<lambda>(x,y). F x y)"
+      apply (rule compact_uniformly_continuous)
+      apply (rule continuous_on_subset[OF cond_uu])
+      using abu \<delta>
+      apply (auto simp: compact_Times simp del: mem_cball)
+      done
+    then obtain \<eta> where "\<eta>>0"
+        and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
+                         dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
+      apply (rule uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"])
+      using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> by auto
+    have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>;   x2 \<in> closed_segment a b;
+              norm (w - x1') \<le> \<delta>;  x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk>
+              \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)"
+             for x1 x2 x1' x2'
+      using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm)
+    have le_ee: "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>"
+                if "x' \<in> u" "cmod (x' - w) < \<delta>" "cmod (x' - w) < \<eta>"  for x'
+    proof -
+      have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)"
+        apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>])
+        apply (rule contour_integrable_diff [OF cont_dw cont_dw])
+        using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> \<open>0 < \<delta>\<close> \<open>w \<in> u\<close> that
+        apply (auto simp: norm_minus_commute)
+        done
+      also have "... = \<epsilon>" using \<open>a \<noteq> b\<close> by simp
+      finally show ?thesis .
+    qed
+    show ?thesis
+      apply (rule_tac x="min \<delta> \<eta>" in exI)
+      using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close>
+      apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> u\<close> intro: le_ee)
+      done
+  qed
+  show ?thesis
+    apply (rule continuous_onI)
+    apply (cases "a=b")
+    apply (auto intro: *)
+    done
+qed
+
+text\<open>This version has @{term"polynomial_function \<gamma>"} as an additional assumption.\<close>
+lemma Cauchy_integral_formula_global_weak:
+    assumes u: "open u" and holf: "f holomorphic_on u"
+        and z: "z \<in> u" and \<gamma>: "polynomial_function \<gamma>"
+        and pasz: "path_image \<gamma> \<subseteq> u - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+        and zero: "\<And>w. w \<notin> u \<Longrightarrow> winding_number \<gamma> w = 0"
+      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
+proof -
+  obtain \<gamma>' where pf\<gamma>': "polynomial_function \<gamma>'" and \<gamma>': "\<And>x. (\<gamma> has_vector_derivative (\<gamma>' x)) (at x)"
+    using has_vector_derivative_polynomial_function [OF \<gamma>] by blast
+  then have "bounded(path_image \<gamma>')"
+    by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function)
+  then obtain B where "B>0" and B: "\<And>x. x \<in> path_image \<gamma>' \<Longrightarrow> norm x \<le> B"
+    using bounded_pos by force
+  define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w
+  define v where "v = {w. w \<notin> path_image \<gamma> \<and> winding_number \<gamma> w = 0}"
+  have "path \<gamma>" "valid_path \<gamma>" using \<gamma>
+    by (auto simp: path_polynomial_function valid_path_polynomial_function)
+  then have ov: "open v"
+    by (simp add: v_def open_winding_number_levelsets loop)
+  have uv_Un: "u \<union> v = UNIV"
+    using pasz zero by (auto simp: v_def)
+  have conf: "continuous_on u f"
+    by (metis holf holomorphic_on_imp_continuous_on)
+  have hol_d: "(d y) holomorphic_on u" if "y \<in> u" for y
+  proof -
+    have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u"
+      by (simp add: holf pole_lemma_open u)
+    then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y"
+      using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce
+    then have "continuous_on u (d y)"
+      apply (simp add: d_def continuous_on_eq_continuous_at u, clarify)
+      using * holomorphic_on_def
+      by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u)
+    moreover have "d y holomorphic_on u - {y}"
+      apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric])
+      apply (intro ballI)
+      apply (rename_tac w)
+      apply (rule_tac d="dist w y" and f = "\<lambda>w. (f w - f y)/(w - y)" in field_differentiable_transform_within)
+      apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros)
+      using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast
+      done
+    ultimately show ?thesis
+      by (rule no_isolated_singularity) (auto simp: u)
+  qed
+  have cint_fxy: "(\<lambda>x. (f x - f y) / (x - y)) contour_integrable_on \<gamma>" if "y \<notin> path_image \<gamma>" for y
+    apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"])
+    using \<open>valid_path \<gamma>\<close> pasz
+    apply (auto simp: u open_delete)
+    apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] |
+                force simp add: that)+
+    done
+  define h where
+    "h z = (if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z)))" for z
+  have U: "\<And>z. z \<in> u \<Longrightarrow> ((d z) has_contour_integral h z) \<gamma>"
+    apply (simp add: h_def)
+    apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
+    using u pasz \<open>valid_path \<gamma>\<close>
+    apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
+    done
+  have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z
+  proof -
+    have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z"
+      using v_def z by auto
+    then have "((\<lambda>x. 1 / (x - z)) has_contour_integral 0) \<gamma>"
+     using z v_def  has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close>] by fastforce
+    then have "((\<lambda>x. f z * (1 / (x - z))) has_contour_integral 0) \<gamma>"
+      using has_contour_integral_lmul by fastforce
+    then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>"
+      by (simp add: divide_simps)
+    moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
+      using z
+      apply (auto simp: v_def)
+      apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy)
+      done
+    ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>"
+      by (rule has_contour_integral_add)
+    have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
+            if  "z \<in> u"
+      using * by (auto simp: divide_simps has_contour_integral_eq)
+    moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>"
+            if "z \<notin> u"
+      apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
+      using u pasz \<open>valid_path \<gamma>\<close> that
+      apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
+      apply (rule continuous_intros conf holomorphic_intros holf | force)+
+      done
+    ultimately show ?thesis
+      using z by (simp add: h_def)
+  qed
+  have znot: "z \<notin> path_image \<gamma>"
+    using pasz by blast
+  obtain d0 where "d0>0" and d0: "\<And>x y. x \<in> path_image \<gamma> \<Longrightarrow> y \<in> - u \<Longrightarrow> d0 \<le> dist x y"
+    using separate_compact_closed [of "path_image \<gamma>" "-u"] pasz u
+    by (fastforce simp add: \<open>path \<gamma>\<close> compact_path_image)
+  obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> u"
+    apply (rule that [of "d0/2"])
+    using \<open>0 < d0\<close>
+    apply (auto simp: dist_norm dest: d0)
+    done
+  have "\<And>x x'. \<lbrakk>x \<in> path_image \<gamma>; dist x x' * 2 < dd\<rbrakk> \<Longrightarrow> \<exists>y k. x' = y + k \<and> y \<in> path_image \<gamma> \<and> dist 0 k * 2 \<le> dd"
+    apply (rule_tac x=x in exI)
+    apply (rule_tac x="x'-x" in exI)
+    apply (force simp add: dist_norm)
+    done
+  then have 1: "path_image \<gamma> \<subseteq> interior {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}"
+    apply (clarsimp simp add: mem_interior)
+    using \<open>0 < dd\<close>
+    apply (rule_tac x="dd/2" in exI, auto)
+    done
+  obtain t where "compact t" and subt: "path_image \<gamma> \<subseteq> interior t" and t: "t \<subseteq> u"
+    apply (rule that [OF _ 1])
+    apply (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums)
+    apply (rule order_trans [OF _ dd])
+    using \<open>0 < dd\<close> by fastforce
+  obtain L where "L>0"
+           and L: "\<And>f B. \<lbrakk>f holomorphic_on interior t; \<And>z. z\<in>interior t \<Longrightarrow> cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
+                         cmod (contour_integral \<gamma> f) \<le> L * B"
+      using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt]
+      by blast
+  have "bounded(f ` t)"
+    by (meson \<open>compact t\<close> compact_continuous_image compact_imp_bounded conf continuous_on_subset t)
+  then obtain D where "D>0" and D: "\<And>x. x \<in> t \<Longrightarrow> norm (f x) \<le> D"
+    by (auto simp: bounded_pos)
+  obtain C where "C>0" and C: "\<And>x. x \<in> t \<Longrightarrow> norm x \<le> C"
+    using \<open>compact t\<close> bounded_pos compact_imp_bounded by force
+  have "dist (h y) 0 \<le> e" if "0 < e" and le: "D * L / e + C \<le> cmod y" for e y
+  proof -
+    have "D * L / e > 0"  using \<open>D>0\<close> \<open>L>0\<close> \<open>e>0\<close> by simp
+    with le have ybig: "norm y > C" by force
+    with C have "y \<notin> t"  by force
+    then have ynot: "y \<notin> path_image \<gamma>"
+      using subt interior_subset by blast
+    have [simp]: "winding_number \<gamma> y = 0"
+      apply (rule winding_number_zero_outside [of _ "cball 0 C"])
+      using ybig interior_subset subt
+      apply (force simp add: loop \<open>path \<gamma>\<close> dist_norm intro!: C)+
+      done
+    have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))"
+      by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V)
+    have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior t"
+      apply (rule holomorphic_on_divide)
+      using holf holomorphic_on_subset interior_subset t apply blast
+      apply (rule holomorphic_intros)+
+      using \<open>y \<notin> t\<close> interior_subset by auto
+    have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior t" for z
+    proof -
+      have "D * L / e + cmod z \<le> cmod y"
+        using le C [of z] z using interior_subset by force
+      then have DL2: "D * L / e \<le> cmod (z - y)"
+        using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute)
+      have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))"
+        by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse)
+      also have "... \<le> D * (e / L / D)"
+        apply (rule mult_mono)
+        using that D interior_subset apply blast
+        using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2
+        apply (auto simp: norm_divide divide_simps algebra_simps)
+        done
+      finally show ?thesis .
+    qed
+    have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))"
+      by (simp add: dist_norm)
+    also have "... \<le> L * (D * (e / L / D))"
+      by (rule L [OF holint leD])
+    also have "... = e"
+      using  \<open>L>0\<close> \<open>0 < D\<close> by auto
+    finally show ?thesis .
+  qed
+  then have "(h \<longlongrightarrow> 0) at_infinity"
+    by (meson Lim_at_infinityI)
+  moreover have "h holomorphic_on UNIV"
+  proof -
+    have con_ff: "continuous (at (x,z)) (\<lambda>(x,y). (f y - f x) / (y - x))"
+                 if "x \<in> u" "z \<in> u" "x \<noteq> z" for x z
+      using that conf
+      apply (simp add: split_def continuous_on_eq_continuous_at u)
+      apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+
+      done
+    have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)"
+      by (rule continuous_intros)+
+    have open_uu_Id: "open (u \<times> u - Id)"
+      apply (rule open_Diff)
+      apply (simp add: open_Times u)
+      using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0]
+      apply (auto simp: Id_fstsnd_eq algebra_simps)
+      done
+    have con_derf: "continuous (at z) (deriv f)" if "z \<in> u" for z
+      apply (rule continuous_on_interior [of u])
+      apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u)
+      by (simp add: interior_open that u)
+    have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x)
+                                else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x)
+                      (at (x, x) within u \<times> u)" if "x \<in> u" for x
+    proof (rule Lim_withinI)
+      fix e::real assume "0 < e"
+      obtain k1 where "k1>0" and k1: "\<And>x'. norm (x' - x) \<le> k1 \<Longrightarrow> norm (deriv f x' - deriv f x) < e"
+        using \<open>0 < e\<close> continuous_within_E [OF con_derf [OF \<open>x \<in> u\<close>]]
+        by (metis UNIV_I dist_norm)
+      obtain k2 where "k2>0" and k2: "ball x k2 \<subseteq> u" by (blast intro: openE [OF u] \<open>x \<in> u\<close>)
+      have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e"
+                    if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2"
+                 for x' z'
+      proof -
+        have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w
+          apply (drule segment_furthest_le [where y=x])
+          by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans)
+        have derf_le: "w \<in> closed_segment x' z' \<Longrightarrow> z' \<noteq> x' \<Longrightarrow> cmod (deriv f w - deriv f x) \<le> e" for w
+          by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans)
+        have f_has_der: "\<And>x. x \<in> u \<Longrightarrow> (f has_field_derivative deriv f x) (at x within u)"
+          by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u)
+        have "closed_segment x' z' \<subseteq> u"
+          by (rule order_trans [OF _ k2]) (simp add: cs_less  le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff)
+        then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')"
+          using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz  by simp
+        then have *: "((\<lambda>x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')"
+          by (rule has_contour_integral_div)
+        have "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e/norm(z' - x') * norm(z' - x')"
+          apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]])
+          using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']]
+                 \<open>e > 0\<close>  \<open>z' \<noteq> x'\<close>
+          apply (auto simp: norm_divide divide_simps derf_le)
+          done
+        also have "... \<le> e" using \<open>0 < e\<close> by simp
+        finally show ?thesis .
+      qed
+      show "\<exists>d>0. \<forall>xa\<in>u \<times> u.
+                  0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow>
+                  dist (case xa of (x, y) \<Rightarrow> if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \<le> e"
+        apply (rule_tac x="min k1 k2" in exI)
+        using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close>
+        apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le)
+        done
+    qed
+    have con_pa_f: "continuous_on (path_image \<gamma>) f"
+      by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t)
+    have le_B: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at t)) \<le> B"
+      apply (rule B)
+      using \<gamma>' using path_image_def vector_derivative_at by fastforce
+    have f_has_cint: "\<And>w. w \<in> v - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f u / (u - w) ^ 1) has_contour_integral h w) \<gamma>"
+      by (simp add: V)
+    have cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). d x y)"
+      apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f')
+      apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify)
+      apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\<lambda>(x,y). (f y - f x) / (y - x))"])
+      using con_ff
+      apply (auto simp: continuous_within)
+      done
+    have hol_dw: "(\<lambda>z. d z w) holomorphic_on u" if "w \<in> u" for w
+    proof -
+      have "continuous_on u ((\<lambda>(x,y). d x y) o (\<lambda>z. (w,z)))"
+        by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+
+      then have *: "continuous_on u (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))"
+        by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps)
+      have **: "\<And>x. \<lbrakk>x \<in> u; x \<noteq> w\<rbrakk> \<Longrightarrow> (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x"
+        apply (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within)
+        apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+
+        done
+      show ?thesis
+        unfolding d_def
+        apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"])
+        apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **)
+        done
+    qed
+    { fix a b
+      assume abu: "closed_segment a b \<subseteq> u"
+      then have "\<And>w. w \<in> u \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)"
+        by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on)
+      then have cont_cint_d: "continuous_on u (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
+        apply (rule contour_integral_continuous_on_linepath_2D [OF \<open>open u\<close> _ _ abu])
+        apply (auto simp: intro: continuous_on_swap_args cond_uu)
+        done
+      have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) o \<gamma>)"
+        apply (rule continuous_on_compose)
+        using \<open>path \<gamma>\<close> path_def pasz
+        apply (auto intro!: continuous_on_subset [OF cont_cint_d])
+        apply (force simp add: path_image_def)
+        done
+      have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
+        apply (simp add: contour_integrable_on)
+        apply (rule integrable_continuous_real)
+        apply (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
+        using pf\<gamma>'
+        by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
+      have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))"
+        using abu  by (force simp add: h_def intro: contour_integral_eq)
+      also have "... =  contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
+        apply (rule contour_integral_swap)
+        apply (rule continuous_on_subset [OF cond_uu])
+        using abu pasz \<open>valid_path \<gamma>\<close>
+        apply (auto intro!: continuous_intros)
+        by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at)
+      finally have cint_h_eq:
+          "contour_integral (linepath a b) h =
+                    contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" .
+      note cint_cint cint_h_eq
+    } note cint_h = this
+    have conthu: "continuous_on u h"
+    proof (simp add: continuous_on_sequentially, clarify)
+      fix a x
+      assume x: "x \<in> u" and au: "\<forall>n. a n \<in> u" and ax: "a \<longlonglongrightarrow> x"
+      then have A1: "\<forall>\<^sub>F n in sequentially. d (a n) contour_integrable_on \<gamma>"
+        by (meson U contour_integrable_on_def eventuallyI)
+      obtain dd where "dd>0" and dd: "cball x dd \<subseteq> u" using open_contains_cball u x by force
+      have A2: "\<forall>\<^sub>F n in sequentially. \<forall>xa\<in>path_image \<gamma>. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee
+      proof -
+        let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
+        have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)"
+          apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]])
+          using dd pasz \<open>valid_path \<gamma>\<close>
+          apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball)
+          done
+        then obtain kk where "kk>0"
+            and kk: "\<And>x x'. \<lbrakk>x\<in>?ddpa; x'\<in>?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow>
+                             dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee"
+          apply (rule uniformly_continuous_onE [where e = ee])
+          using \<open>0 < ee\<close> by auto
+        have kk: "\<lbrakk>norm (w - x) \<le> dd; z \<in> path_image \<gamma>; norm ((w, z) - (x, z)) < kk\<rbrakk> \<Longrightarrow> norm (d w z - d x z) < ee"
+                 for  w z
+          using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm)
+        show ?thesis
+          using ax unfolding lim_sequentially eventually_sequentially
+          apply (drule_tac x="min dd kk" in spec)
+          using \<open>dd > 0\<close> \<open>kk > 0\<close>
+          apply (fastforce simp: kk dist_norm)
+          done
+      qed
+      have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x"
+        apply (simp add: contour_integral_unique [OF U, symmetric] x)
+        apply (rule contour_integral_uniform_limit [OF A1 A2 le_B])
+        apply (auto simp: \<open>valid_path \<gamma>\<close>)
+        done
+      then show "(h \<circ> a) \<longlonglongrightarrow> h x"
+        by (simp add: h_def x au o_def)
+    qed
+    show ?thesis
+    proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify)
+      fix z0
+      consider "z0 \<in> v" | "z0 \<in> u" using uv_Un by blast
+      then show "h field_differentiable at z0"
+      proof cases
+        assume "z0 \<in> v" then show ?thesis
+          using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \<open>valid_path \<gamma>\<close>
+          by (auto simp: field_differentiable_def v_def)
+      next
+        assume "z0 \<in> u" then
+        obtain e where "e>0" and e: "ball z0 e \<subseteq> u" by (blast intro: openE [OF u])
+        have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0"
+                if abc_subset: "convex hull {a, b, c} \<subseteq> ball z0 e"  for a b c
+        proof -
+          have *: "\<And>x1 x2 z. z \<in> u \<Longrightarrow> closed_segment x1 x2 \<subseteq> u \<Longrightarrow> (\<lambda>w. d w z) contour_integrable_on linepath x1 x2"
+            using  hol_dw holomorphic_on_imp_continuous_on u
+            by (auto intro!: contour_integrable_holomorphic_simple)
+          have abc: "closed_segment a b \<subseteq> u"  "closed_segment b c \<subseteq> u"  "closed_segment c a \<subseteq> u"
+            using that e segments_subset_convex_hull by fastforce+
+          have eq0: "\<And>w. w \<in> u \<Longrightarrow> contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\<lambda>z. d z w) = 0"
+            apply (rule contour_integral_unique [OF Cauchy_theorem_triangle])
+            apply (rule holomorphic_on_subset [OF hol_dw])
+            using e abc_subset by auto
+          have "contour_integral \<gamma>
+                   (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) +
+                        (contour_integral (linepath b c) (\<lambda>z. d z x) +
+                         contour_integral (linepath c a) (\<lambda>z. d z x)))  =  0"
+            apply (rule contour_integral_eq_0)
+            using abc pasz u
+            apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+
+            done
+          then show ?thesis
+            by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac)
+        qed
+        show ?thesis
+          using e \<open>e > 0\<close>
+          by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic
+                           Morera_triangle continuous_on_subset [OF conthu] *)
+      qed
+    qed
+  qed
+  ultimately have [simp]: "h z = 0" for z
+    by (meson Liouville_weak)
+  have "((\<lambda>w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z) \<gamma>"
+    by (rule has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close> znot])
+  then have "((\<lambda>w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
+    by (metis mult.commute has_contour_integral_lmul)
+  then have 1: "((\<lambda>w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
+    by (simp add: divide_simps)
+  moreover have 2: "((\<lambda>w. (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
+    using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\<lambda>w. (f w - f z)/(w - z)"])
+  show ?thesis
+    using has_contour_integral_add [OF 1 2]  by (simp add: diff_divide_distrib)
+qed
+
+
+theorem Cauchy_integral_formula_global:
+    assumes s: "open s" and holf: "f holomorphic_on s"
+        and z: "z \<in> s" and vpg: "valid_path \<gamma>"
+        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+        and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
+      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
+proof -
+  have "path \<gamma>" using vpg by (blast intro: valid_path_imp_path)
+  have hols: "(\<lambda>w. f w / (w - z)) holomorphic_on s - {z}" "(\<lambda>w. 1 / (w - z)) holomorphic_on s - {z}"
+    by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+
+  then have cint_fw: "(\<lambda>w. f w / (w - z)) contour_integrable_on \<gamma>"
+    by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz)
+  obtain d where "d>0"
+      and d: "\<And>g h. \<lbrakk>valid_path g; valid_path h; \<forall>t\<in>{0..1}. cmod (g t - \<gamma> t) < d \<and> cmod (h t - \<gamma> t) < d;
+                     pathstart h = pathstart g \<and> pathfinish h = pathfinish g\<rbrakk>
+                     \<Longrightarrow> path_image h \<subseteq> s - {z} \<and> (\<forall>f. f holomorphic_on s - {z} \<longrightarrow> contour_integral h f = contour_integral g f)"
+    using contour_integral_nearby_ends [OF _ \<open>path \<gamma>\<close> pasz] s by (simp add: open_Diff) metis
+  obtain p where polyp: "polynomial_function p"
+             and ps: "pathstart p = pathstart \<gamma>" and pf: "pathfinish p = pathfinish \<gamma>" and led: "\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < d"
+    using path_approx_polynomial_function [OF \<open>path \<gamma>\<close> \<open>d > 0\<close>] by blast
+  then have ploop: "pathfinish p = pathstart p" using loop by auto
+  have vpp: "valid_path p"  using polyp valid_path_polynomial_function by blast
+  have [simp]: "z \<notin> path_image \<gamma>" using pasz by blast
+  have paps: "path_image p \<subseteq> s - {z}" and cint_eq: "(\<And>f. f holomorphic_on s - {z} \<Longrightarrow> contour_integral p f = contour_integral \<gamma> f)"
+    using pf ps led d [OF vpg vpp] \<open>d > 0\<close> by auto
+  have wn_eq: "winding_number p z = winding_number \<gamma> z"
+    using vpp paps
+    by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols)
+  have "winding_number p w = winding_number \<gamma> w" if "w \<notin> s" for w
+  proof -
+    have hol: "(\<lambda>v. 1 / (v - w)) holomorphic_on s - {z}"
+      using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf])
+   have "w \<notin> path_image p" "w \<notin> path_image \<gamma>" using paps pasz that by auto
+   then show ?thesis
+    using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol])
+  qed
+  then have wn0: "\<And>w. w \<notin> s \<Longrightarrow> winding_number p w = 0"
+    by (simp add: zero)
+  show ?thesis
+    using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols
+    by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq)
+qed
+
+theorem Cauchy_theorem_global:
+    assumes s: "open s" and holf: "f holomorphic_on s"
+        and vpg: "valid_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
+        and pas: "path_image \<gamma> \<subseteq> s"
+        and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
+      shows "(f has_contour_integral 0) \<gamma>"
+proof -
+  obtain z where "z \<in> s" and znot: "z \<notin> path_image \<gamma>"
+  proof -
+    have "compact (path_image \<gamma>)"
+      using compact_valid_path_image vpg by blast
+    then have "path_image \<gamma> \<noteq> s"
+      by (metis (no_types) compact_open path_image_nonempty s)
+    with pas show ?thesis by (blast intro: that)
+  qed
+  then have pasz: "path_image \<gamma> \<subseteq> s - {z}" using pas by blast
+  have hol: "(\<lambda>w. (w - z) * f w) holomorphic_on s"
+    by (rule holomorphic_intros holf)+
+  show ?thesis
+    using Cauchy_integral_formula_global [OF s hol \<open>z \<in> s\<close> vpg pasz loop zero]
+    by (auto simp: znot elim!: has_contour_integral_eq)
+qed
+
+corollary Cauchy_theorem_global_outside:
+    assumes "open s" "f holomorphic_on s" "valid_path \<gamma>"  "pathfinish \<gamma> = pathstart \<gamma>" "path_image \<gamma> \<subseteq> s"
+            "\<And>w. w \<notin> s \<Longrightarrow> w \<in> outside(path_image \<gamma>)"
+      shows "(f has_contour_integral 0) \<gamma>"
+by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path)
+
+
+end
--- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,7539 +0,0 @@
-section \<open>Complex path integrals and Cauchy's integral theorem\<close>
-
-text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
-
-theory Cauchy_Integral_Thm
-imports Complex_Transcendental Weierstrass Ordered_Euclidean_Space
-begin
-
-subsection\<open>Homeomorphisms of arc images\<close>
-
-lemma homeomorphism_arc:
-  fixes g :: "real \<Rightarrow> 'a::t2_space"
-  assumes "arc g"
-  obtains h where "homeomorphism {0..1} (path_image g) g h"
-using assms by (force simp add: arc_def homeomorphism_compact path_def path_image_def)
-
-lemma homeomorphic_arc_image_interval:
-  fixes g :: "real \<Rightarrow> 'a::t2_space" and a::real
-  assumes "arc g" "a < b"
-  shows "(path_image g) homeomorphic {a..b}"
-proof -
-  have "(path_image g) homeomorphic {0..1::real}"
-    by (meson assms(1) homeomorphic_def homeomorphic_sym homeomorphism_arc)
-  also have "... homeomorphic {a..b}"
-    using assms by (force intro: homeomorphic_closed_intervals_real)
-  finally show ?thesis .
-qed
-
-lemma homeomorphic_arc_images:
-  fixes g :: "real \<Rightarrow> 'a::t2_space" and h :: "real \<Rightarrow> 'b::t2_space"
-  assumes "arc g" "arc h"
-  shows "(path_image g) homeomorphic (path_image h)"
-proof -
-  have "(path_image g) homeomorphic {0..1::real}"
-    by (meson assms homeomorphic_def homeomorphic_sym homeomorphism_arc)
-  also have "... homeomorphic (path_image h)"
-    by (meson assms homeomorphic_def homeomorphism_arc)
-  finally show ?thesis .
-qed
-
-subsection \<open>Piecewise differentiable functions\<close>
-
-definition piecewise_differentiable_on
-           (infixr "piecewise'_differentiable'_on" 50)
-  where "f piecewise_differentiable_on i  \<equiv>
-           continuous_on i f \<and>
-           (\<exists>s. finite s \<and> (\<forall>x \<in> i - s. f differentiable (at x within i)))"
-
-lemma piecewise_differentiable_on_imp_continuous_on:
-    "f piecewise_differentiable_on s \<Longrightarrow> continuous_on s f"
-by (simp add: piecewise_differentiable_on_def)
-
-lemma piecewise_differentiable_on_subset:
-    "f piecewise_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_differentiable_on t"
-  using continuous_on_subset
-  unfolding piecewise_differentiable_on_def
-  apply safe
-  apply (blast intro: elim: continuous_on_subset)
-  by (meson Diff_iff differentiable_within_subset subsetCE)
-
-lemma differentiable_on_imp_piecewise_differentiable:
-  fixes a:: "'a::{linorder_topology,real_normed_vector}"
-  shows "f differentiable_on {a..b} \<Longrightarrow> f piecewise_differentiable_on {a..b}"
-  apply (simp add: piecewise_differentiable_on_def differentiable_imp_continuous_on)
-  apply (rule_tac x="{a,b}" in exI, simp add: differentiable_on_def)
-  done
-
-lemma differentiable_imp_piecewise_differentiable:
-    "(\<And>x. x \<in> s \<Longrightarrow> f differentiable (at x within s))
-         \<Longrightarrow> f piecewise_differentiable_on s"
-by (auto simp: piecewise_differentiable_on_def differentiable_imp_continuous_on differentiable_on_def
-         intro: differentiable_within_subset)
-
-lemma piecewise_differentiable_const [iff]: "(\<lambda>x. z) piecewise_differentiable_on s"
-  by (simp add: differentiable_imp_piecewise_differentiable)
-
-lemma piecewise_differentiable_compose:
-    "\<lbrakk>f piecewise_differentiable_on s; g piecewise_differentiable_on (f ` s);
-      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
-      \<Longrightarrow> (g o f) piecewise_differentiable_on s"
-  apply (simp add: piecewise_differentiable_on_def, safe)
-  apply (blast intro: continuous_on_compose2)
-  apply (rename_tac A B)
-  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
-  apply (blast intro: differentiable_chain_within)
-  done
-
-lemma piecewise_differentiable_affine:
-  fixes m::real
-  assumes "f piecewise_differentiable_on ((\<lambda>x. m *\<^sub>R x + c) ` s)"
-  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_differentiable_on s"
-proof (cases "m = 0")
-  case True
-  then show ?thesis
-    unfolding o_def
-    by (force intro: differentiable_imp_piecewise_differentiable differentiable_const)
-next
-  case False
-  show ?thesis
-    apply (rule piecewise_differentiable_compose [OF differentiable_imp_piecewise_differentiable])
-    apply (rule assms derivative_intros | simp add: False vimage_def real_vector_affinity_eq)+
-    done
-qed
-
-lemma piecewise_differentiable_cases:
-  fixes c::real
-  assumes "f piecewise_differentiable_on {a..c}"
-          "g piecewise_differentiable_on {c..b}"
-           "a \<le> c" "c \<le> b" "f c = g c"
-  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_differentiable_on {a..b}"
-proof -
-  obtain s t where st: "finite s" "finite t"
-                       "\<forall>x\<in>{a..c} - s. f differentiable at x within {a..c}"
-                       "\<forall>x\<in>{c..b} - t. g differentiable at x within {c..b}"
-    using assms
-    by (auto simp: piecewise_differentiable_on_def)
-  have finabc: "finite ({a,b,c} \<union> (s \<union> t))"
-    by (metis \<open>finite s\<close> \<open>finite t\<close> finite_Un finite_insert finite.emptyI)
-  have "continuous_on {a..c} f" "continuous_on {c..b} g"
-    using assms piecewise_differentiable_on_def by auto
-  then have "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
-    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
-                               OF closed_real_atLeastAtMost [of c b],
-                               of f g "\<lambda>x. x\<le>c"]  assms
-    by (force simp: ivl_disj_un_two_touch)
-  moreover
-  { fix x
-    assume x: "x \<in> {a..b} - ({a,b,c} \<union> (s \<union> t))"
-    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b}" (is "?diff_fg")
-    proof (cases x c rule: le_cases)
-      case le show ?diff_fg
-        apply (rule differentiable_transform_within [where d = "dist x c" and f = f])
-        using x le st
-        apply (simp_all add: dist_real_def)
-        apply (rule differentiable_at_withinI)
-        apply (rule differentiable_within_open [where s = "{a<..<c} - s", THEN iffD1], simp_all)
-        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
-        apply (force elim!: differentiable_subset)+
-        done
-    next
-      case ge show ?diff_fg
-        apply (rule differentiable_transform_within [where d = "dist x c" and f = g])
-        using x ge st
-        apply (simp_all add: dist_real_def)
-        apply (rule differentiable_at_withinI)
-        apply (rule differentiable_within_open [where s = "{c<..<b} - t", THEN iffD1], simp_all)
-        apply (blast intro: open_greaterThanLessThan finite_imp_closed)
-        apply (force elim!: differentiable_subset)+
-        done
-    qed
-  }
-  then have "\<exists>s. finite s \<and>
-                 (\<forall>x\<in>{a..b} - s. (\<lambda>x. if x \<le> c then f x else g x) differentiable at x within {a..b})"
-    by (meson finabc)
-  ultimately show ?thesis
-    by (simp add: piecewise_differentiable_on_def)
-qed
-
-lemma piecewise_differentiable_neg:
-    "f piecewise_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_differentiable_on s"
-  by (auto simp: piecewise_differentiable_on_def continuous_on_minus)
-
-lemma piecewise_differentiable_add:
-  assumes "f piecewise_differentiable_on i"
-          "g piecewise_differentiable_on i"
-    shows "(\<lambda>x. f x + g x) piecewise_differentiable_on i"
-proof -
-  obtain s t where st: "finite s" "finite t"
-                       "\<forall>x\<in>i - s. f differentiable at x within i"
-                       "\<forall>x\<in>i - t. g differentiable at x within i"
-    using assms by (auto simp: piecewise_differentiable_on_def)
-  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)"
-    by auto
-  moreover have "continuous_on i f" "continuous_on i g"
-    using assms piecewise_differentiable_on_def by auto
-  ultimately show ?thesis
-    by (auto simp: piecewise_differentiable_on_def continuous_on_add)
-qed
-
-lemma piecewise_differentiable_diff:
-    "\<lbrakk>f piecewise_differentiable_on s;  g piecewise_differentiable_on s\<rbrakk>
-     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_differentiable_on s"
-  unfolding diff_conv_add_uminus
-  by (metis piecewise_differentiable_add piecewise_differentiable_neg)
-
-lemma continuous_on_joinpaths_D1:
-    "continuous_on {0..1} (g1 +++ g2) \<Longrightarrow> continuous_on {0..1} g1"
-  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (op*(inverse 2))"])
-  apply (rule continuous_intros | simp)+
-  apply (auto elim!: continuous_on_subset simp: joinpaths_def)
-  done
-
-lemma continuous_on_joinpaths_D2:
-    "\<lbrakk>continuous_on {0..1} (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> continuous_on {0..1} g2"
-  apply (rule continuous_on_eq [of _ "(g1 +++ g2) o (\<lambda>x. inverse 2*x + 1/2)"])
-  apply (rule continuous_intros | simp)+
-  apply (auto elim!: continuous_on_subset simp add: joinpaths_def pathfinish_def pathstart_def Ball_def)
-  done
-
-lemma piecewise_differentiable_D1:
-    "(g1 +++ g2) piecewise_differentiable_on {0..1} \<Longrightarrow> g1 piecewise_differentiable_on {0..1}"
-  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D1)
-  apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
-  apply simp
-  apply (intro ballI)
-  apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))"
-       in differentiable_transform_within)
-  apply (auto simp: dist_real_def joinpaths_def)
-  apply (rule differentiable_chain_within derivative_intros | simp)+
-  apply (rule differentiable_subset)
-  apply (force simp:)+
-  done
-
-lemma piecewise_differentiable_D2:
-    "\<lbrakk>(g1 +++ g2) piecewise_differentiable_on {0..1}; pathfinish g1 = pathstart g2\<rbrakk>
-    \<Longrightarrow> g2 piecewise_differentiable_on {0..1}"
-  apply (clarsimp simp add: piecewise_differentiable_on_def dest!: continuous_on_joinpaths_D2)
-  apply (rule_tac x="insert 0 ((\<lambda>x. 2*x-1)`s)" in exI)
-  apply simp
-  apply (intro ballI)
-  apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)"
-          in differentiable_transform_within)
-  apply (auto simp: dist_real_def joinpaths_def abs_if field_simps split: if_split_asm)
-  apply (rule differentiable_chain_within derivative_intros | simp)+
-  apply (rule differentiable_subset)
-  apply (force simp: divide_simps)+
-  done
-
-
-subsubsection\<open>The concept of continuously differentiable\<close>
-
-text \<open>
-John Harrison writes as follows:
-
-``The usual assumption in complex analysis texts is that a path \<open>\<gamma>\<close> should be piecewise
-continuously differentiable, which ensures that the path integral exists at least for any continuous
-f, since all piecewise continuous functions are integrable. However, our notion of validity is
-weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a
-finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to
-the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this
-can integrate all derivatives.''
-
-"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec.
-Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165.
-
-And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably
-difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem
-asserting that all derivatives can be integrated before we can adopt Harrison's choice.\<close>
-
-definition C1_differentiable_on :: "(real \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> real set \<Rightarrow> bool"
-           (infix "C1'_differentiable'_on" 50)
-  where
-  "f C1_differentiable_on s \<longleftrightarrow>
-   (\<exists>D. (\<forall>x \<in> s. (f has_vector_derivative (D x)) (at x)) \<and> continuous_on s D)"
-
-lemma C1_differentiable_on_eq:
-    "f C1_differentiable_on s \<longleftrightarrow>
-     (\<forall>x \<in> s. f differentiable at x) \<and> continuous_on s (\<lambda>x. vector_derivative f (at x))"
-  unfolding C1_differentiable_on_def
-  apply safe
-  using differentiable_def has_vector_derivative_def apply blast
-  apply (erule continuous_on_eq)
-  using vector_derivative_at apply fastforce
-  using vector_derivative_works apply fastforce
-  done
-
-lemma C1_differentiable_on_subset:
-  "f C1_differentiable_on t \<Longrightarrow> s \<subseteq> t \<Longrightarrow> f C1_differentiable_on s"
-  unfolding C1_differentiable_on_def  continuous_on_eq_continuous_within
-  by (blast intro:  continuous_within_subset)
-
-lemma C1_differentiable_compose:
-    "\<lbrakk>f C1_differentiable_on s; g C1_differentiable_on (f ` s);
-      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
-      \<Longrightarrow> (g o f) C1_differentiable_on s"
-  apply (simp add: C1_differentiable_on_eq, safe)
-   using differentiable_chain_at apply blast
-  apply (rule continuous_on_eq [of _ "\<lambda>x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"])
-   apply (rule Limits.continuous_on_scaleR, assumption)
-   apply (metis (mono_tags, lifting) continuous_on_eq continuous_at_imp_continuous_on continuous_on_compose differentiable_imp_continuous_within o_def)
-  by (simp add: vector_derivative_chain_at)
-
-lemma C1_diff_imp_diff: "f C1_differentiable_on s \<Longrightarrow> f differentiable_on s"
-  by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on)
-
-lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\<lambda>x. x) C1_differentiable_on s"
-  by (auto simp: C1_differentiable_on_eq continuous_on_const)
-
-lemma C1_differentiable_on_const [simp, derivative_intros]: "(\<lambda>z. a) C1_differentiable_on s"
-  by (auto simp: C1_differentiable_on_eq continuous_on_const)
-
-lemma C1_differentiable_on_add [simp, derivative_intros]:
-  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x + g x) C1_differentiable_on s"
-  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
-
-lemma C1_differentiable_on_minus [simp, derivative_intros]:
-  "f C1_differentiable_on s \<Longrightarrow> (\<lambda>x. - f x) C1_differentiable_on s"
-  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
-
-lemma C1_differentiable_on_diff [simp, derivative_intros]:
-  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x - g x) C1_differentiable_on s"
-  unfolding C1_differentiable_on_eq  by (auto intro: continuous_intros)
-
-lemma C1_differentiable_on_mult [simp, derivative_intros]:
-  fixes f g :: "real \<Rightarrow> 'a :: real_normed_algebra"
-  shows "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x * g x) C1_differentiable_on s"
-  unfolding C1_differentiable_on_eq
-  by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within)
-
-lemma C1_differentiable_on_scaleR [simp, derivative_intros]:
-  "f C1_differentiable_on s \<Longrightarrow> g C1_differentiable_on s \<Longrightarrow> (\<lambda>x. f x *\<^sub>R g x) C1_differentiable_on s"
-  unfolding C1_differentiable_on_eq
-  by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+
-
-
-definition piecewise_C1_differentiable_on
-           (infixr "piecewise'_C1'_differentiable'_on" 50)
-  where "f piecewise_C1_differentiable_on i  \<equiv>
-           continuous_on i f \<and>
-           (\<exists>s. finite s \<and> (f C1_differentiable_on (i - s)))"
-
-lemma C1_differentiable_imp_piecewise:
-    "f C1_differentiable_on s \<Longrightarrow> f piecewise_C1_differentiable_on s"
-  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within)
-
-lemma piecewise_C1_imp_differentiable:
-    "f piecewise_C1_differentiable_on i \<Longrightarrow> f piecewise_differentiable_on i"
-  by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def
-           C1_differentiable_on_def differentiable_def has_vector_derivative_def
-           intro: has_derivative_at_within)
-
-lemma piecewise_C1_differentiable_compose:
-    "\<lbrakk>f piecewise_C1_differentiable_on s; g piecewise_C1_differentiable_on (f ` s);
-      \<And>x. finite (s \<inter> f-`{x})\<rbrakk>
-      \<Longrightarrow> (g o f) piecewise_C1_differentiable_on s"
-  apply (simp add: piecewise_C1_differentiable_on_def, safe)
-  apply (blast intro: continuous_on_compose2)
-  apply (rename_tac A B)
-  apply (rule_tac x="A \<union> (\<Union>x\<in>B. s \<inter> f-`{x})" in exI)
-  apply (rule conjI, blast)
-  apply (rule C1_differentiable_compose)
-  apply (blast intro: C1_differentiable_on_subset)
-  apply (blast intro: C1_differentiable_on_subset)
-  by (simp add: Diff_Int_distrib2)
-
-lemma piecewise_C1_differentiable_on_subset:
-    "f piecewise_C1_differentiable_on s \<Longrightarrow> t \<le> s \<Longrightarrow> f piecewise_C1_differentiable_on t"
-  by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset)
-
-lemma C1_differentiable_imp_continuous_on:
-  "f C1_differentiable_on s \<Longrightarrow> continuous_on s f"
-  unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within
-  using differentiable_at_withinI differentiable_imp_continuous_within by blast
-
-lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}"
-  unfolding C1_differentiable_on_def
-  by auto
-
-lemma piecewise_C1_differentiable_affine:
-  fixes m::real
-  assumes "f piecewise_C1_differentiable_on ((\<lambda>x. m * x + c) ` s)"
-  shows "(f o (\<lambda>x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on s"
-proof (cases "m = 0")
-  case True
-  then show ?thesis
-    unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const)
-next
-  case False
-  show ?thesis
-    apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise])
-    apply (rule assms derivative_intros | simp add: False vimage_def)+
-    using real_vector_affinity_eq [OF False, where c=c, unfolded scaleR_conv_of_real]
-    apply simp
-    done
-qed
-
-lemma piecewise_C1_differentiable_cases:
-  fixes c::real
-  assumes "f piecewise_C1_differentiable_on {a..c}"
-          "g piecewise_C1_differentiable_on {c..b}"
-           "a \<le> c" "c \<le> b" "f c = g c"
-  shows "(\<lambda>x. if x \<le> c then f x else g x) piecewise_C1_differentiable_on {a..b}"
-proof -
-  obtain s t where st: "f C1_differentiable_on ({a..c} - s)"
-                       "g C1_differentiable_on ({c..b} - t)"
-                       "finite s" "finite t"
-    using assms
-    by (force simp: piecewise_C1_differentiable_on_def)
-  then have f_diff: "f differentiable_on {a..<c} - s"
-        and g_diff: "g differentiable_on {c<..b} - t"
-    by (simp_all add: C1_differentiable_on_eq differentiable_at_withinI differentiable_on_def)
-  have "continuous_on {a..c} f" "continuous_on {c..b} g"
-    using assms piecewise_C1_differentiable_on_def by auto
-  then have cab: "continuous_on {a..b} (\<lambda>x. if x \<le> c then f x else g x)"
-    using continuous_on_cases [OF closed_real_atLeastAtMost [of a c],
-                               OF closed_real_atLeastAtMost [of c b],
-                               of f g "\<lambda>x. x\<le>c"]  assms
-    by (force simp: ivl_disj_un_two_touch)
-  { fix x
-    assume x: "x \<in> {a..b} - insert c (s \<union> t)"
-    have "(\<lambda>x. if x \<le> c then f x else g x) differentiable at x" (is "?diff_fg")
-    proof (cases x c rule: le_cases)
-      case le show ?diff_fg
-        apply (rule differentiable_transform_within [where f=f and d = "dist x c"])
-        using x dist_real_def le st by (auto simp: C1_differentiable_on_eq)
-    next
-      case ge show ?diff_fg
-        apply (rule differentiable_transform_within [where f=g and d = "dist x c"])
-        using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq)
-    qed
-  }
-  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)"
-    by auto
-  moreover
-  { assume fcon: "continuous_on ({a<..<c} - s) (\<lambda>x. vector_derivative f (at x))"
-       and gcon: "continuous_on ({c<..<b} - t) (\<lambda>x. vector_derivative g (at x))"
-    have "open ({a<..<c} - s)"  "open ({c<..<b} - t)"
-      using st by (simp_all add: open_Diff finite_imp_closed)
-    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))"
-      apply (rule continuous_on_eq [OF fcon])
-      apply (simp add:)
-      apply (rule vector_derivative_at [symmetric])
-      apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within)
-      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
-      apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff linorder_not_le order_less_irrefl st(1))
-      apply auto
-      done
-    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))"
-      apply (rule continuous_on_eq [OF gcon])
-      apply (simp add:)
-      apply (rule vector_derivative_at [symmetric])
-      apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within)
-      apply (simp_all add: dist_norm vector_derivative_works [symmetric])
-      apply (metis (full_types) C1_differentiable_on_eq Diff_iff Groups.add_ac(2) add_mono_thms_linordered_field(5) atLeastAtMost_iff less_irrefl not_le st(2))
-      apply auto
-      done
-    ultimately 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))"
-      apply (rule continuous_on_subset [OF continuous_on_open_Un], auto)
-      done
-  } note * = this
-  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))"
-    using st
-    by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *)
-  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)"
-    apply (rule_tac x="{a,b,c} \<union> s \<union> t" in exI)
-    using st  by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset)
-  with cab show ?thesis
-    by (simp add: piecewise_C1_differentiable_on_def)
-qed
-
-lemma piecewise_C1_differentiable_neg:
-    "f piecewise_C1_differentiable_on s \<Longrightarrow> (\<lambda>x. -(f x)) piecewise_C1_differentiable_on s"
-  unfolding piecewise_C1_differentiable_on_def
-  by (auto intro!: continuous_on_minus C1_differentiable_on_minus)
-
-lemma piecewise_C1_differentiable_add:
-  assumes "f piecewise_C1_differentiable_on i"
-          "g piecewise_C1_differentiable_on i"
-    shows "(\<lambda>x. f x + g x) piecewise_C1_differentiable_on i"
-proof -
-  obtain s t where st: "finite s" "finite t"
-                       "f C1_differentiable_on (i-s)"
-                       "g C1_differentiable_on (i-t)"
-    using assms by (auto simp: piecewise_C1_differentiable_on_def)
-  then have "finite (s \<union> t) \<and> (\<lambda>x. f x + g x) C1_differentiable_on i - (s \<union> t)"
-    by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset)
-  moreover have "continuous_on i f" "continuous_on i g"
-    using assms piecewise_C1_differentiable_on_def by auto
-  ultimately show ?thesis
-    by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add)
-qed
-
-lemma piecewise_C1_differentiable_diff:
-    "\<lbrakk>f piecewise_C1_differentiable_on s;  g piecewise_C1_differentiable_on s\<rbrakk>
-     \<Longrightarrow> (\<lambda>x. f x - g x) piecewise_C1_differentiable_on s"
-  unfolding diff_conv_add_uminus
-  by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg)
-
-lemma piecewise_C1_differentiable_D1:
-  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
-  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}"
-    shows "g1 piecewise_C1_differentiable_on {0..1}"
-proof -
-  obtain s where "finite s"
-             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
-             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
-    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  then have g1D: "g1 differentiable at x" if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
-    apply (rule_tac d="dist (x/2) (1/2)" and f = "(g1 +++ g2) o (op*(inverse 2))" in differentiable_transform_within)
-    using that
-    apply (simp_all add: dist_real_def joinpaths_def)
-    apply (rule differentiable_chain_at derivative_intros | force)+
-    done
-  have [simp]: "vector_derivative (g1 \<circ> op * 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)"
-               if "x \<in> {0..1} - insert 1 (op * 2 ` s)" for x
-    apply (subst vector_derivative_chain_at)
-    using that
-    apply (rule derivative_eq_intros g1D | simp)+
-    done
-  have "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
-    using co12 by (rule continuous_on_subset) force
-  then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 o op*2) (at x))"
-    apply (rule continuous_on_eq [OF _ vector_derivative_at])
-    apply (rule_tac f="g1 o op*2" and d="dist x (1/2)" in has_vector_derivative_transform_within)
-    apply (simp_all add: dist_norm joinpaths_def vector_derivative_works [symmetric])
-    apply (force intro: g1D differentiable_chain_at)
-    apply auto
-    done
-  have "continuous_on ({0..1} - insert 1 (op * 2 ` s))
-                      ((\<lambda>x. 1/2 * vector_derivative (g1 o op*2) (at x)) o op*(1/2))"
-    apply (rule continuous_intros)+
-    using coDhalf
-    apply (simp add: scaleR_conv_of_real image_set_diff image_image)
-    done
-  then have con_g1: "continuous_on ({0..1} - insert 1 (op * 2 ` s)) (\<lambda>x. vector_derivative g1 (at x))"
-    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
-  have "continuous_on {0..1} g1"
-    using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast
-  with \<open>finite s\<close> show ?thesis
-    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-    apply (rule_tac x="insert 1 ((op*2)`s)" in exI)
-    apply (simp add: g1D con_g1)
-  done
-qed
-
-lemma piecewise_C1_differentiable_D2:
-  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
-  assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2"
-    shows "g2 piecewise_C1_differentiable_on {0..1}"
-proof -
-  obtain s where "finite s"
-             and co12: "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
-             and g12D: "\<forall>x\<in>{0..1} - s. g1 +++ g2 differentiable at x"
-    using assms  by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  then have g2D: "g2 differentiable at x" if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
-    apply (rule_tac d="dist ((x+1)/2) (1/2)" and f = "(g1 +++ g2) o (\<lambda>x. (x+1)/2)" in differentiable_transform_within)
-    using that
-    apply (simp_all add: dist_real_def joinpaths_def)
-    apply (auto simp: dist_real_def joinpaths_def field_simps)
-    apply (rule differentiable_chain_at derivative_intros | force)+
-    apply (drule_tac x= "(x + 1) / 2" in bspec, force simp: divide_simps)
-    apply assumption
-    done
-  have [simp]: "vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)"
-               if "x \<in> {0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)" for x
-    using that  by (auto simp: vector_derivative_chain_at divide_simps g2D)
-  have "continuous_on ({1/2..1} - insert (1/2) s) (\<lambda>x. vector_derivative (g1 +++ g2) (at x))"
-    using co12 by (rule continuous_on_subset) force
-  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))"
-    apply (rule continuous_on_eq [OF _ vector_derivative_at])
-    apply (rule_tac f="g2 o (\<lambda>x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within)
-    apply (auto simp: dist_real_def field_simps joinpaths_def vector_derivative_works [symmetric]
-                intro!: g2D differentiable_chain_at)
-    done
-  have [simp]: "((\<lambda>x. (x + 1) / 2) ` ({0..1} - insert 0 ((\<lambda>x. 2 * x - 1) ` s))) = ({1/2..1} - insert (1/2) s)"
-    apply (simp add: image_set_diff inj_on_def image_image)
-    apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib)
-    done
-  have "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s))
-                      ((\<lambda>x. 1/2 * vector_derivative (g2 \<circ> (\<lambda>x. 2*x-1)) (at x)) o (\<lambda>x. (x+1)/2))"
-    by (rule continuous_intros | simp add:  coDhalf)+
-  then have con_g2: "continuous_on ({0..1} - insert 0 ((\<lambda>x. 2*x-1) ` s)) (\<lambda>x. vector_derivative g2 (at x))"
-    by (rule continuous_on_eq) (simp add: scaleR_conv_of_real)
-  have "continuous_on {0..1} g2"
-    using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast
-  with \<open>finite s\<close> show ?thesis
-    apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-    apply (rule_tac x="insert 0 ((\<lambda>x. 2 * x - 1) ` s)" in exI)
-    apply (simp add: g2D con_g2)
-  done
-qed
-
-subsection \<open>Valid paths, and their start and finish\<close>
-
-lemma Diff_Un_eq: "A - (B \<union> C) = A - B - C"
-  by blast
-
-definition valid_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
-  where "valid_path f \<equiv> f piecewise_C1_differentiable_on {0..1::real}"
-
-definition closed_path :: "(real \<Rightarrow> 'a :: real_normed_vector) \<Rightarrow> bool"
-  where "closed_path g \<equiv> g 0 = g 1"
-
-subsubsection\<open>In particular, all results for paths apply\<close>
-
-lemma valid_path_imp_path: "valid_path g \<Longrightarrow> path g"
-by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def)
-
-lemma connected_valid_path_image: "valid_path g \<Longrightarrow> connected(path_image g)"
-  by (metis connected_path_image valid_path_imp_path)
-
-lemma compact_valid_path_image: "valid_path g \<Longrightarrow> compact(path_image g)"
-  by (metis compact_path_image valid_path_imp_path)
-
-lemma bounded_valid_path_image: "valid_path g \<Longrightarrow> bounded(path_image g)"
-  by (metis bounded_path_image valid_path_imp_path)
-
-lemma closed_valid_path_image: "valid_path g \<Longrightarrow> closed(path_image g)"
-  by (metis closed_path_image valid_path_imp_path)
-
-proposition valid_path_compose:
-  assumes "valid_path g"
-      and der: "\<And>x. x \<in> path_image g \<Longrightarrow> \<exists>f'. (f has_field_derivative f') (at x)"
-      and con: "continuous_on (path_image g) (deriv f)"
-    shows "valid_path (f o g)"
-proof -
-  obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s"
-    using \<open>valid_path g\<close> unfolding valid_path_def piecewise_C1_differentiable_on_def by auto
-  have "f \<circ> g differentiable at t" when "t\<in>{0..1} - s" for t
-    proof (rule differentiable_chain_at)
-      show "g differentiable at t" using \<open>valid_path g\<close>
-        by (meson C1_differentiable_on_eq \<open>g C1_differentiable_on {0..1} - s\<close> that)
-    next
-      have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
-      then obtain f' where "(f has_field_derivative f') (at (g t))"
-        using der by auto
-      then have " (f has_derivative op * f') (at (g t))"
-        using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto
-      then show "f differentiable at (g t)" using differentiableI by auto
-    qed
-  moreover have "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative (f \<circ> g) (at x))"
-    proof (rule continuous_on_eq [where f = "\<lambda>x. vector_derivative g (at x) * deriv f (g x)"],
-        rule continuous_intros)
-      show "continuous_on ({0..1} - s) (\<lambda>x. vector_derivative g (at x))"
-        using g_diff C1_differentiable_on_eq by auto
-    next
-      have "continuous_on {0..1} (\<lambda>x. deriv f (g x))"
-        using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def]
-          \<open>valid_path g\<close> piecewise_C1_differentiable_on_def valid_path_def
-        by blast
-      then show "continuous_on ({0..1} - s) (\<lambda>x. deriv f (g x))"
-        using continuous_on_subset by blast
-    next
-      show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \<circ> g) (at t)"
-          when "t \<in> {0..1} - s" for t
-        proof (rule vector_derivative_chain_at_general[symmetric])
-          show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that)
-        next
-          have "g t\<in>path_image g" using that DiffD1 image_eqI path_image_def by metis
-          then obtain f' where "(f has_field_derivative f') (at (g t))"
-            using der by auto
-          then show "\<exists>g'. (f has_field_derivative g') (at (g t))" by auto
-        qed
-    qed
-  ultimately have "f o g C1_differentiable_on {0..1} - s"
-    using C1_differentiable_on_eq by blast
-  moreover have "path (f o g)"
-    proof -
-      have "isCont f x" when "x\<in>path_image g" for x
-        proof -
-          obtain f' where "(f has_field_derivative f') (at x)"
-            using der[rule_format] \<open>x\<in>path_image g\<close> by auto
-          thus ?thesis using DERIV_isCont by auto
-        qed
-      then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto
-      then show ?thesis using path_continuous_image \<open>valid_path g\<close> valid_path_imp_path by auto
-    qed
-  ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def
-    using \<open>finite s\<close> by auto
-qed
-
-
-subsection\<open>Contour Integrals along a path\<close>
-
-text\<open>This definition is for complex numbers only, and does not generalise to line integrals in a vector field\<close>
-
-text\<open>piecewise differentiable function on [0,1]\<close>
-
-definition has_contour_integral :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool"
-           (infixr "has'_contour'_integral" 50)
-  where "(f has_contour_integral i) g \<equiv>
-           ((\<lambda>x. f(g x) * vector_derivative g (at x within {0..1}))
-            has_integral i) {0..1}"
-
-definition contour_integrable_on
-           (infixr "contour'_integrable'_on" 50)
-  where "f contour_integrable_on g \<equiv> \<exists>i. (f has_contour_integral i) g"
-
-definition contour_integral
-  where "contour_integral g f \<equiv> @i. (f has_contour_integral i) g \<or> ~ f contour_integrable_on g \<and> i=0"
-
-lemma not_integrable_contour_integral: "~ f contour_integrable_on g \<Longrightarrow> contour_integral g f = 0"
-  unfolding contour_integrable_on_def contour_integral_def by blast
-
-lemma contour_integral_unique: "(f has_contour_integral i) g \<Longrightarrow> contour_integral g f = i"
-  apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def)
-  using has_integral_unique by blast
-
-corollary has_contour_integral_eqpath:
-     "\<lbrakk>(f has_contour_integral y) p; f contour_integrable_on \<gamma>;
-       contour_integral p f = contour_integral \<gamma> f\<rbrakk>
-      \<Longrightarrow> (f has_contour_integral y) \<gamma>"
-using contour_integrable_on_def contour_integral_unique by auto
-
-lemma has_contour_integral_integral:
-    "f contour_integrable_on i \<Longrightarrow> (f has_contour_integral (contour_integral i f)) i"
-  by (metis contour_integral_unique contour_integrable_on_def)
-
-lemma has_contour_integral_unique:
-    "(f has_contour_integral i) g \<Longrightarrow> (f has_contour_integral j) g \<Longrightarrow> i = j"
-  using has_integral_unique
-  by (auto simp: has_contour_integral_def)
-
-lemma has_contour_integral_integrable: "(f has_contour_integral i) g \<Longrightarrow> f contour_integrable_on g"
-  using contour_integrable_on_def by blast
-
-(* Show that we can forget about the localized derivative.*)
-
-lemma vector_derivative_within_interior:
-     "\<lbrakk>x \<in> interior s; NO_MATCH UNIV s\<rbrakk>
-      \<Longrightarrow> vector_derivative f (at x within s) = vector_derivative f (at x)"
-  apply (simp add: vector_derivative_def has_vector_derivative_def has_derivative_def netlimit_within_interior)
-  apply (subst lim_within_interior, auto)
-  done
-
-lemma has_integral_localized_vector_derivative:
-    "((\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \<longleftrightarrow>
-     ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}"
-proof -
-  have "{a..b} - {a,b} = interior {a..b}"
-    by (simp add: atLeastAtMost_diff_ends)
-  show ?thesis
-    apply (rule has_integral_spike_eq [of "{a,b}"])
-    apply (auto simp: vector_derivative_within_interior)
-    done
-qed
-
-lemma integrable_on_localized_vector_derivative:
-    "(\<lambda>x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \<longleftrightarrow>
-     (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}"
-  by (simp add: integrable_on_def has_integral_localized_vector_derivative)
-
-lemma has_contour_integral:
-     "(f has_contour_integral i) g \<longleftrightarrow>
-      ((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
-  by (simp add: has_integral_localized_vector_derivative has_contour_integral_def)
-
-lemma contour_integrable_on:
-     "f contour_integrable_on g \<longleftrightarrow>
-      (\<lambda>t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}"
-  by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def)
-
-subsection\<open>Reversing a path\<close>
-
-lemma valid_path_imp_reverse:
-  assumes "valid_path g"
-    shows "valid_path(reversepath g)"
-proof -
-  obtain s where "finite s" "g C1_differentiable_on ({0..1} - s)"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
-  then have "finite (op - 1 ` s)" "(reversepath g C1_differentiable_on ({0..1} - op - 1 ` s))"
-    apply (auto simp: reversepath_def)
-    apply (rule C1_differentiable_compose [of "\<lambda>x::real. 1-x" _ g, unfolded o_def])
-    apply (auto simp: C1_differentiable_on_eq)
-    apply (rule continuous_intros, force)
-    apply (force elim!: continuous_on_subset)
-    apply (simp add: finite_vimageI inj_on_def)
-    done
-  then show ?thesis using assms
-    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric])
-qed
-
-lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \<longleftrightarrow> valid_path g"
-  using valid_path_imp_reverse by force
-
-lemma has_contour_integral_reversepath:
-  assumes "valid_path g" "(f has_contour_integral i) g"
-    shows "(f has_contour_integral (-i)) (reversepath g)"
-proof -
-  { fix s x
-    assume xs: "g C1_differentiable_on ({0..1} - s)" "x \<notin> op - 1 ` s" "0 \<le> x" "x \<le> 1"
-      have "vector_derivative (\<lambda>x. g (1 - x)) (at x within {0..1}) =
-            - vector_derivative g (at (1 - x) within {0..1})"
-      proof -
-        obtain f' where f': "(g has_vector_derivative f') (at (1 - x))"
-          using xs
-          by (force simp: has_vector_derivative_def C1_differentiable_on_def)
-        have "(g o (\<lambda>x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)"
-          apply (rule vector_diff_chain_within)
-          apply (intro vector_diff_chain_within derivative_eq_intros | simp)+
-          apply (rule has_vector_derivative_at_within [OF f'])
-          done
-        then have mf': "((\<lambda>x. g (1 - x)) has_vector_derivative -f') (at x)"
-          by (simp add: o_def)
-        show ?thesis
-          using xs
-          by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f'])
-      qed
-  } note * = this
-  have 01: "{0..1::real} = cbox 0 1"
-    by simp
-  show ?thesis using assms
-    apply (auto simp: has_contour_integral_def)
-    apply (drule has_integral_affinity01 [where m= "-1" and c=1])
-    apply (auto simp: reversepath_def valid_path_def piecewise_C1_differentiable_on_def)
-    apply (drule has_integral_neg)
-    apply (rule_tac s = "(\<lambda>x. 1 - x) ` s" in has_integral_spike_finite)
-    apply (auto simp: *)
-    done
-qed
-
-lemma contour_integrable_reversepath:
-    "valid_path g \<Longrightarrow> f contour_integrable_on g \<Longrightarrow> f contour_integrable_on (reversepath g)"
-  using has_contour_integral_reversepath contour_integrable_on_def by blast
-
-lemma contour_integrable_reversepath_eq:
-    "valid_path g \<Longrightarrow> (f contour_integrable_on (reversepath g) \<longleftrightarrow> f contour_integrable_on g)"
-  using contour_integrable_reversepath valid_path_reversepath by fastforce
-
-lemma contour_integral_reversepath:
-  assumes "valid_path g"
-    shows "contour_integral (reversepath g) f = - (contour_integral g f)"
-proof (cases "f contour_integrable_on g")
-  case True then show ?thesis
-    by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath)
-next
-  case False then have "~ f contour_integrable_on (reversepath g)"
-    by (simp add: assms contour_integrable_reversepath_eq)
-  with False show ?thesis by (simp add: not_integrable_contour_integral)
-qed
-
-
-subsection\<open>Joining two paths together\<close>
-
-lemma valid_path_join:
-  assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2"
-    shows "valid_path(g1 +++ g2)"
-proof -
-  have "g1 1 = g2 0"
-    using assms by (auto simp: pathfinish_def pathstart_def)
-  moreover have "(g1 o (\<lambda>x. 2*x)) piecewise_C1_differentiable_on {0..1/2}"
-    apply (rule piecewise_C1_differentiable_compose)
-    using assms
-    apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths)
-    apply (rule continuous_intros | simp)+
-    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
-    done
-  moreover have "(g2 o (\<lambda>x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}"
-    apply (rule piecewise_C1_differentiable_compose)
-    using assms unfolding valid_path_def piecewise_C1_differentiable_on_def
-    by (auto intro!: continuous_intros finite_vimageI [where h = "(\<lambda>x. 2*x - 1)"] inj_onI
-             simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths)
-  ultimately show ?thesis
-    apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def)
-    apply (rule piecewise_C1_differentiable_cases)
-    apply (auto simp: o_def)
-    done
-qed
-
-lemma valid_path_join_D1:
-  fixes g1 :: "real \<Rightarrow> 'a::real_normed_field"
-  shows "valid_path (g1 +++ g2) \<Longrightarrow> valid_path g1"
-  unfolding valid_path_def
-  by (rule piecewise_C1_differentiable_D1)
-
-lemma valid_path_join_D2:
-  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
-  shows "\<lbrakk>valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> valid_path g2"
-  unfolding valid_path_def
-  by (rule piecewise_C1_differentiable_D2)
-
-lemma valid_path_join_eq [simp]:
-  fixes g2 :: "real \<Rightarrow> 'a::real_normed_field"
-  shows "pathfinish g1 = pathstart g2 \<Longrightarrow> (valid_path(g1 +++ g2) \<longleftrightarrow> valid_path g1 \<and> valid_path g2)"
-  using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast
-
-lemma has_contour_integral_join:
-  assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2"
-          "valid_path g1" "valid_path g2"
-    shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)"
-proof -
-  obtain s1 s2
-    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
-      and s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
-    using assms
-    by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  have 1: "((\<lambda>x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}"
-   and 2: "((\<lambda>x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}"
-    using assms
-    by (auto simp: has_contour_integral)
-  have i1: "((\<lambda>x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}"
-   and i2: "((\<lambda>x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}"
-    using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]]
-          has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]]
-    by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac)
-  have g1: "\<lbrakk>0 \<le> z; z*2 < 1; z*2 \<notin> s1\<rbrakk> \<Longrightarrow>
-            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
-            2 *\<^sub>R vector_derivative g1 (at (z*2))" for z
-    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>z - 1/2\<bar>"]])
-    apply (simp_all add: dist_real_def abs_if split: if_split_asm)
-    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x" 2 _ g1, simplified o_def])
-    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
-    using s1
-    apply (auto simp: algebra_simps vector_derivative_works)
-    done
-  have g2: "\<lbrakk>1 < z*2; z \<le> 1; z*2 - 1 \<notin> s2\<rbrakk> \<Longrightarrow>
-            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at z) =
-            2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z
-    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2 (2*x - 1))" and d = "\<bar>z - 1/2\<bar>"]])
-    apply (simp_all add: dist_real_def abs_if split: if_split_asm)
-    apply (rule vector_diff_chain_at [of "\<lambda>x. 2*x - 1" 2 _ g2, simplified o_def])
-    apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
-    using s2
-    apply (auto simp: algebra_simps vector_derivative_works)
-    done
-  have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}"
-    apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) (op*2 -` s1)"])
-    using s1
-    apply (force intro: finite_vimageI [where h = "op*2"] inj_onI)
-    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1)
-    done
-  moreover have "((\<lambda>x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}"
-    apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\<lambda>x. 2*x-1) -` s2)"])
-    using s2
-    apply (force intro: finite_vimageI [where h = "\<lambda>x. 2*x-1"] inj_onI)
-    apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2)
-    done
-  ultimately
-  show ?thesis
-    apply (simp add: has_contour_integral)
-    apply (rule has_integral_combine [where c = "1/2"], auto)
-    done
-qed
-
-lemma contour_integrable_joinI:
-  assumes "f contour_integrable_on g1" "f contour_integrable_on g2"
-          "valid_path g1" "valid_path g2"
-    shows "f contour_integrable_on (g1 +++ g2)"
-  using assms
-  by (meson has_contour_integral_join contour_integrable_on_def)
-
-lemma contour_integrable_joinD1:
-  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1"
-    shows "f contour_integrable_on g1"
-proof -
-  obtain s1
-    where s1: "finite s1" "\<forall>x\<in>{0..1} - s1. g1 differentiable at x"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  have "(\<lambda>x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
-    using assms
-    apply (auto simp: contour_integrable_on)
-    apply (drule integrable_on_subcbox [where a=0 and b="1/2"])
-    apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified])
-    done
-  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}"
-    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
-  have g1: "\<lbrakk>0 < z; z < 1; z \<notin> s1\<rbrakk> \<Longrightarrow>
-            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) =
-            2 *\<^sub>R vector_derivative g1 (at z)"  for z
-    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g1(2*x))" and d = "\<bar>(z-1)/2\<bar>"]])
-    apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
-    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2" 2 _ g1, simplified o_def])
-    using s1
-    apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left)
-    done
-  show ?thesis
-    using s1
-    apply (auto simp: contour_integrable_on)
-    apply (rule integrable_spike_finite [of "{0,1} \<union> s1", OF _ _ *])
-    apply (auto simp: joinpaths_def scaleR_conv_of_real g1)
-    done
-qed
-
-lemma contour_integrable_joinD2:
-  assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2"
-    shows "f contour_integrable_on g2"
-proof -
-  obtain s2
-    where s2: "finite s2" "\<forall>x\<in>{0..1} - s2. g2 differentiable at x"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  have "(\<lambda>x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}"
-    using assms
-    apply (auto simp: contour_integrable_on)
-    apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto)
-    apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified])
-    apply (simp add: image_affinity_atLeastAtMost_diff)
-    done
-  then have *: "(\<lambda>x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2)))
-                integrable_on {0..1}"
-    by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real)
-  have g2: "\<lbrakk>0 < z; z < 1; z \<notin> s2\<rbrakk> \<Longrightarrow>
-            vector_derivative (\<lambda>x. if x*2 \<le> 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) =
-            2 *\<^sub>R vector_derivative g2 (at z)" for z
-    apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g2(2*x-1))" and d = "\<bar>z/2\<bar>"]])
-    apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm)
-    apply (rule vector_diff_chain_at [of "\<lambda>x. x*2-1" 2 _ g2, simplified o_def])
-    using s2
-    apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left
-                      vector_derivative_works add_divide_distrib)
-    done
-  show ?thesis
-    using s2
-    apply (auto simp: contour_integrable_on)
-    apply (rule integrable_spike_finite [of "{0,1} \<union> s2", OF _ _ *])
-    apply (auto simp: joinpaths_def scaleR_conv_of_real g2)
-    done
-qed
-
-lemma contour_integrable_join [simp]:
-  shows
-    "\<lbrakk>valid_path g1; valid_path g2\<rbrakk>
-     \<Longrightarrow> f contour_integrable_on (g1 +++ g2) \<longleftrightarrow> f contour_integrable_on g1 \<and> f contour_integrable_on g2"
-using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast
-
-lemma contour_integral_join [simp]:
-  shows
-    "\<lbrakk>f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\<rbrakk>
-        \<Longrightarrow> contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f"
-  by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique)
-
-
-subsection\<open>Shifting the starting point of a (closed) path\<close>
-
-lemma shiftpath_alt_def: "shiftpath a f = (\<lambda>x. if x \<le> 1-a then f (a + x) else f (a + x - 1))"
-  by (auto simp: shiftpath_def)
-
-lemma valid_path_shiftpath [intro]:
-  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "valid_path(shiftpath a g)"
-  using assms
-  apply (auto simp: valid_path_def shiftpath_alt_def)
-  apply (rule piecewise_C1_differentiable_cases)
-  apply (auto simp: algebra_simps)
-  apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one])
-  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
-  apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps])
-  apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset)
-  done
-
-lemma has_contour_integral_shiftpath:
-  assumes f: "(f has_contour_integral i) g" "valid_path g"
-      and a: "a \<in> {0..1}"
-    shows "(f has_contour_integral i) (shiftpath a g)"
-proof -
-  obtain s
-    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  have *: "((\<lambda>x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}"
-    using assms by (auto simp: has_contour_integral)
-  then have i: "i = integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)) +
-                    integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x))"
-    apply (rule has_integral_unique)
-    apply (subst add.commute)
-    apply (subst Integration.integral_combine)
-    using assms * integral_unique by auto
-  { fix x
-    have "0 \<le> x \<Longrightarrow> x + a < 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a) ` s \<Longrightarrow>
-         vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))"
-      unfolding shiftpath_def
-      apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x))" and d = "dist(1-a) x"]])
-        apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
-      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a" 1 _ g, simplified o_def scaleR_one])
-       apply (intro derivative_eq_intros | simp)+
-      using g
-       apply (drule_tac x="x+a" in bspec)
-      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
-      done
-  } note vd1 = this
-  { fix x
-    have "1 < x + a \<Longrightarrow> x \<le> 1 \<Longrightarrow> x \<notin> (\<lambda>x. x - a + 1) ` s \<Longrightarrow>
-          vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))"
-      unfolding shiftpath_def
-      apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\<lambda>x. g(a+x-1))" and d = "dist (1-a) x"]])
-        apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm)
-      apply (rule vector_diff_chain_at [of "\<lambda>x. x+a-1" 1 _ g, simplified o_def scaleR_one])
-       apply (intro derivative_eq_intros | simp)+
-      using g
-      apply (drule_tac x="x+a-1" in bspec)
-      using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute)
-      done
-  } note vd2 = this
-  have va1: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})"
-    using * a   by (fastforce intro: integrable_subinterval_real)
-  have v0a: "(\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})"
-    apply (rule integrable_subinterval_real)
-    using * a by auto
-  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
-        has_integral  integral {a..1} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {0..1 - a}"
-    apply (rule has_integral_spike_finite
-             [where s = "{1-a} \<union> (\<lambda>x. x-a) ` s" and f = "\<lambda>x. f(g(a+x)) * vector_derivative g (at(a+x))"])
-      using s apply blast
-     using a apply (auto simp: algebra_simps vd1)
-     apply (force simp: shiftpath_def add.commute)
-    using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]]
-    apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute)
-    done
-  moreover
-  have "((\<lambda>x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x))
-        has_integral  integral {0..a} (\<lambda>x. f (g x) * vector_derivative g (at x)))  {1 - a..1}"
-    apply (rule has_integral_spike_finite
-             [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))"])
-      using s apply blast
-     using a apply (auto simp: algebra_simps vd2)
-     apply (force simp: shiftpath_def add.commute)
-    using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]]
-    apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified])
-    apply (simp add: algebra_simps)
-    done
-  ultimately show ?thesis
-    using a
-    by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"])
-qed
-
-lemma has_contour_integral_shiftpath_D:
-  assumes "(f has_contour_integral i) (shiftpath a g)"
-          "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "(f has_contour_integral i) g"
-proof -
-  obtain s
-    where s: "finite s" and g: "\<forall>x\<in>{0..1} - s. g differentiable at x"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  { fix x
-    assume x: "0 < x" "x < 1" "x \<notin> s"
-    then have gx: "g differentiable at x"
-      using g by auto
-    have "vector_derivative g (at x within {0..1}) =
-          vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})"
-      apply (rule vector_derivative_at_within_ivl
-                  [OF has_vector_derivative_transform_within_open
-                      [where f = "(shiftpath (1 - a) (shiftpath a g))" and s = "{0<..<1}-s"]])
-      using s g assms x
-      apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath
-                        vector_derivative_within_interior vector_derivative_works [symmetric])
-      apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"])
-      apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm)
-      done
-  } note vd = this
-  have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))"
-    using assms  by (auto intro!: has_contour_integral_shiftpath)
-  show ?thesis
-    apply (simp add: has_contour_integral_def)
-    apply (rule has_integral_spike_finite [of "{0,1} \<union> s", OF _ _  fi [unfolded has_contour_integral_def]])
-    using s assms vd
-    apply (auto simp: Path_Connected.shiftpath_shiftpath)
-    done
-qed
-
-lemma has_contour_integral_shiftpath_eq:
-  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "(f has_contour_integral i) (shiftpath a g) \<longleftrightarrow> (f has_contour_integral i) g"
-  using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast
-
-lemma contour_integrable_on_shiftpath_eq:
-  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "f contour_integrable_on (shiftpath a g) \<longleftrightarrow> f contour_integrable_on g"
-using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto
-
-lemma contour_integral_shiftpath:
-  assumes "valid_path g" "pathfinish g = pathstart g" "a \<in> {0..1}"
-    shows "contour_integral (shiftpath a g) f = contour_integral g f"
-   using assms
-   by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq)
-
-
-subsection\<open>More about straight-line paths\<close>
-
-lemma has_vector_derivative_linepath_within:
-    "(linepath a b has_vector_derivative (b - a)) (at x within s)"
-apply (simp add: linepath_def has_vector_derivative_def algebra_simps)
-apply (rule derivative_eq_intros | simp)+
-done
-
-lemma vector_derivative_linepath_within:
-    "x \<in> {0..1} \<Longrightarrow> vector_derivative (linepath a b) (at x within {0..1}) = b - a"
-  apply (rule vector_derivative_within_closed_interval [of 0 "1::real", simplified])
-  apply (auto simp: has_vector_derivative_linepath_within)
-  done
-
-lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a"
-  by (simp add: has_vector_derivative_linepath_within vector_derivative_at)
-
-lemma valid_path_linepath [iff]: "valid_path (linepath a b)"
-  apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath)
-  apply (rule_tac x="{}" in exI)
-  apply (simp add: differentiable_on_def differentiable_def)
-  using has_vector_derivative_def has_vector_derivative_linepath_within
-  apply (fastforce simp add: continuous_on_eq_continuous_within)
-  done
-
-lemma has_contour_integral_linepath:
-  shows "(f has_contour_integral i) (linepath a b) \<longleftrightarrow>
-         ((\<lambda>x. f(linepath a b x) * (b - a)) has_integral i) {0..1}"
-  by (simp add: has_contour_integral vector_derivative_linepath_at)
-
-lemma linepath_in_path:
-  shows "x \<in> {0..1} \<Longrightarrow> linepath a b x \<in> closed_segment a b"
-  by (auto simp: segment linepath_def)
-
-lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b"
-  by (auto simp: segment linepath_def)
-
-lemma linepath_in_convex_hull:
-    fixes x::real
-    assumes a: "a \<in> convex hull s"
-        and b: "b \<in> convex hull s"
-        and x: "0\<le>x" "x\<le>1"
-       shows "linepath a b x \<in> convex hull s"
-  apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD])
-  using x
-  apply (auto simp: linepath_image_01 [symmetric])
-  done
-
-lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b"
-  by (simp add: linepath_def)
-
-lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0"
-  by (simp add: linepath_def)
-
-lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)"
-  by (simp add: has_contour_integral_linepath)
-
-lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0"
-  using has_contour_integral_trivial contour_integral_unique by blast
-
-
-subsection\<open>Relation to subpath construction\<close>
-
-lemma valid_path_subpath:
-  fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
-  assumes "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
-    shows "valid_path(subpath u v g)"
-proof (cases "v=u")
-  case True
-  then show ?thesis
-    unfolding valid_path_def subpath_def
-    by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise)
-next
-  case False
-  have "(g o (\<lambda>x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}"
-    apply (rule piecewise_C1_differentiable_compose)
-    apply (simp add: C1_differentiable_imp_piecewise)
-     apply (simp add: image_affinity_atLeastAtMost)
-    using assms False
-    apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset)
-    apply (subst Int_commute)
-    apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI)
-    done
-  then show ?thesis
-    by (auto simp: o_def valid_path_def subpath_def)
-qed
-
-lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)"
-  by (simp add: has_contour_integral subpath_def)
-
-lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)"
-  using has_contour_integral_subpath_refl contour_integrable_on_def by blast
-
-lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0"
-  by (simp add: has_contour_integral_subpath_refl contour_integral_unique)
-
-lemma has_contour_integral_subpath:
-  assumes f: "f contour_integrable_on g" and g: "valid_path g"
-      and uv: "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
-    shows "(f has_contour_integral  integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x)))
-           (subpath u v g)"
-proof (cases "v=u")
-  case True
-  then show ?thesis
-    using f   by (simp add: contour_integrable_on_def subpath_def has_contour_integral)
-next
-  case False
-  obtain s where s: "\<And>x. x \<in> {0..1} - s \<Longrightarrow> g differentiable at x" and fs: "finite s"
-    using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast
-  have *: "((\<lambda>x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u)))
-            has_integral (1 / (v - u)) * integral {u..v} (\<lambda>t. f (g t) * vector_derivative g (at t)))
-           {0..1}"
-    using f uv
-    apply (simp add: contour_integrable_on subpath_def has_contour_integral)
-    apply (drule integrable_on_subcbox [where a=u and b=v, simplified])
-    apply (simp_all add: has_integral_integral)
-    apply (drule has_integral_affinity [where m="v-u" and c=u, simplified])
-    apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real)
-    apply (simp add: divide_simps False)
-    done
-  { fix x
-    have "x \<in> {0..1} \<Longrightarrow>
-           x \<notin> (\<lambda>t. (v-u) *\<^sub>R t + u) -` s \<Longrightarrow>
-           vector_derivative (\<lambda>x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))"
-      apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]])
-      apply (intro derivative_eq_intros | simp)+
-      apply (cut_tac s [of "(v - u) * x + u"])
-      using uv mult_left_le [of x "v-u"]
-      apply (auto simp:  vector_derivative_works)
-      done
-  } note vd = this
-  show ?thesis
-    apply (cut_tac has_integral_cmul [OF *, where c = "v-u"])
-    using fs assms
-    apply (simp add: False subpath_def has_contour_integral)
-    apply (rule_tac s = "(\<lambda>t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite)
-    apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real)
-    done
-qed
-
-lemma contour_integrable_subpath:
-  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}"
-    shows "f contour_integrable_on (subpath u v g)"
-  apply (cases u v rule: linorder_class.le_cases)
-   apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms])
-  apply (subst reversepath_subpath [symmetric])
-  apply (rule contour_integrable_reversepath)
-   using assms apply (blast intro: valid_path_subpath)
-  apply (simp add: contour_integrable_on_def)
-  using assms apply (blast intro: has_contour_integral_subpath)
-  done
-
-lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i"
-  by blast
-
-lemma has_integral_contour_integral_subpath:
-  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
-    shows "(((\<lambda>x. f(g x) * vector_derivative g (at x)))
-            has_integral  contour_integral (subpath u v g) f) {u..v}"
-  using assms
-  apply (auto simp: has_integral_integrable_integral)
-  apply (rule integrable_on_subcbox [where a=u and b=v and s = "{0..1}", simplified])
-  apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on)
-  done
-
-lemma contour_integral_subcontour_integral:
-  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<le> v"
-    shows "contour_integral (subpath u v g) f =
-           integral {u..v} (\<lambda>x. f(g x) * vector_derivative g (at x))"
-  using assms has_contour_integral_subpath contour_integral_unique by blast
-
-lemma contour_integral_subpath_combine_less:
-  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
-          "u<v" "v<w"
-    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
-           contour_integral (subpath u w g) f"
-  using assms apply (auto simp: contour_integral_subcontour_integral)
-  apply (rule integral_combine, auto)
-  apply (rule integrable_on_subcbox [where a=u and b=w and s = "{0..1}", simplified])
-  apply (auto simp: contour_integrable_on)
-  done
-
-lemma contour_integral_subpath_combine:
-  assumes "f contour_integrable_on g" "valid_path g" "u \<in> {0..1}" "v \<in> {0..1}" "w \<in> {0..1}"
-    shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f =
-           contour_integral (subpath u w g) f"
-proof (cases "u\<noteq>v \<and> v\<noteq>w \<and> u\<noteq>w")
-  case True
-    have *: "subpath v u g = reversepath(subpath u v g) \<and>
-             subpath w u g = reversepath(subpath u w g) \<and>
-             subpath w v g = reversepath(subpath v w g)"
-      by (auto simp: reversepath_subpath)
-    have "u < v \<and> v < w \<or>
-          u < w \<and> w < v \<or>
-          v < u \<and> u < w \<or>
-          v < w \<and> w < u \<or>
-          w < u \<and> u < v \<or>
-          w < v \<and> v < u"
-      using True assms by linarith
-    with assms show ?thesis
-      using contour_integral_subpath_combine_less [of f g u v w]
-            contour_integral_subpath_combine_less [of f g u w v]
-            contour_integral_subpath_combine_less [of f g v u w]
-            contour_integral_subpath_combine_less [of f g v w u]
-            contour_integral_subpath_combine_less [of f g w u v]
-            contour_integral_subpath_combine_less [of f g w v u]
-      apply simp
-      apply (elim disjE)
-      apply (auto simp: * contour_integral_reversepath contour_integrable_subpath
-                   valid_path_reversepath valid_path_subpath algebra_simps)
-      done
-next
-  case False
-  then show ?thesis
-    apply (auto simp: contour_integral_subpath_refl)
-    using assms
-    by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath)
-qed
-
-lemma contour_integral_integral:
-     "contour_integral g f = integral {0..1} (\<lambda>x. f (g x) * vector_derivative g (at x))"
-  by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on)
-
-
-text\<open>Cauchy's theorem where there's a primitive\<close>
-
-lemma contour_integral_primitive_lemma:
-  fixes f :: "complex \<Rightarrow> complex" and g :: "real \<Rightarrow> complex"
-  assumes "a \<le> b"
-      and "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
-      and "g piecewise_differentiable_on {a..b}"  "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
-    shows "((\<lambda>x. f'(g x) * vector_derivative g (at x within {a..b}))
-             has_integral (f(g b) - f(g a))) {a..b}"
-proof -
-  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"
-    using assms by (auto simp: piecewise_differentiable_on_def)
-  have cfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
-    apply (rule continuous_on_compose [OF cg, unfolded o_def])
-    using assms
-    apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff)
-    done
-  { fix x::real
-    assume a: "a < x" and b: "x < b" and xk: "x \<notin> k"
-    then have "g differentiable at x within {a..b}"
-      using k by (simp add: differentiable_at_withinI)
-    then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})"
-      by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real)
-    then have gdiff: "(g has_derivative (\<lambda>u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})"
-      by (simp add: has_vector_derivative_def scaleR_conv_of_real)
-    have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})"
-      using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def)
-    then have fdiff: "(f has_derivative op * (f' (g x))) (at (g x) within g ` {a..b})"
-      by (simp add: has_field_derivative_def)
-    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})"
-      using diff_chain_within [OF gdiff fdiff]
-      by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac)
-  } note * = this
-  show ?thesis
-    apply (rule fundamental_theorem_of_calculus_interior_strong)
-    using k assms cfg *
-    apply (auto simp: at_within_closed_interval)
-    done
-qed
-
-lemma contour_integral_primitive:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
-      and "valid_path g" "path_image g \<subseteq> s"
-    shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g"
-  using assms
-  apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def)
-  apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s])
-  done
-
-corollary Cauchy_theorem_primitive:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s)"
-      and "valid_path g"  "path_image g \<subseteq> s" "pathfinish g = pathstart g"
-    shows "(f' has_contour_integral 0) g"
-  using assms
-  by (metis diff_self contour_integral_primitive)
-
-
-text\<open>Existence of path integral for continuous function\<close>
-lemma contour_integrable_continuous_linepath:
-  assumes "continuous_on (closed_segment a b) f"
-  shows "f contour_integrable_on (linepath a b)"
-proof -
-  have "continuous_on {0..1} ((\<lambda>x. f x * (b - a)) o linepath a b)"
-    apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01)
-    apply (rule continuous_intros | simp add: assms)+
-    done
-  then show ?thesis
-    apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric])
-    apply (rule integrable_continuous [of 0 "1::real", simplified])
-    apply (rule continuous_on_eq [where f = "\<lambda>x. f(linepath a b x)*(b - a)"])
-    apply (auto simp: vector_derivative_linepath_within)
-    done
-qed
-
-lemma has_field_der_id: "((\<lambda>x. x\<^sup>2 / 2) has_field_derivative x) (at x)"
-  by (rule has_derivative_imp_has_field_derivative)
-     (rule derivative_intros | simp)+
-
-lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\<lambda>y. y) = (b^2 - a^2)/2"
-  apply (rule contour_integral_unique)
-  using contour_integral_primitive [of UNIV "\<lambda>x. x^2/2" "\<lambda>x. x" "linepath a b"]
-  apply (auto simp: field_simps has_field_der_id)
-  done
-
-lemma contour_integrable_on_const [iff]: "(\<lambda>x. c) contour_integrable_on (linepath a b)"
-  by (simp add: continuous_on_const contour_integrable_continuous_linepath)
-
-lemma contour_integrable_on_id [iff]: "(\<lambda>x. x) contour_integrable_on (linepath a b)"
-  by (simp add: continuous_on_id contour_integrable_continuous_linepath)
-
-
-subsection\<open>Arithmetical combining theorems\<close>
-
-lemma has_contour_integral_neg:
-    "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. -(f x)) has_contour_integral (-i)) g"
-  by (simp add: has_integral_neg has_contour_integral_def)
-
-lemma has_contour_integral_add:
-    "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
-     \<Longrightarrow> ((\<lambda>x. f1 x + f2 x) has_contour_integral (i1 + i2)) g"
-  by (simp add: has_integral_add has_contour_integral_def algebra_simps)
-
-lemma has_contour_integral_diff:
-  "\<lbrakk>(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\<rbrakk>
-         \<Longrightarrow> ((\<lambda>x. f1 x - f2 x) has_contour_integral (i1 - i2)) g"
-  by (simp add: has_integral_sub has_contour_integral_def algebra_simps)
-
-lemma has_contour_integral_lmul:
-  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. c * (f x)) has_contour_integral (c*i)) g"
-apply (simp add: has_contour_integral_def)
-apply (drule has_integral_mult_right)
-apply (simp add: algebra_simps)
-done
-
-lemma has_contour_integral_rmul:
-  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. (f x) * c) has_contour_integral (i*c)) g"
-apply (drule has_contour_integral_lmul)
-apply (simp add: mult.commute)
-done
-
-lemma has_contour_integral_div:
-  "(f has_contour_integral i) g \<Longrightarrow> ((\<lambda>x. f x/c) has_contour_integral (i/c)) g"
-  by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul)
-
-lemma has_contour_integral_eq:
-    "\<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"
-apply (simp add: path_image_def has_contour_integral_def)
-by (metis (no_types, lifting) image_eqI has_integral_eq)
-
-lemma has_contour_integral_bound_linepath:
-  assumes "(f has_contour_integral i) (linepath a b)"
-          "0 \<le> B" "\<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B"
-    shows "norm i \<le> B * norm(b - a)"
-proof -
-  { fix x::real
-    assume x: "0 \<le> x" "x \<le> 1"
-  have "norm (f (linepath a b x)) *
-        norm (vector_derivative (linepath a b) (at x within {0..1})) \<le> B * norm (b - a)"
-    by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x)
-  } note * = this
-  have "norm i \<le> (B * norm (b - a)) * content (cbox 0 (1::real))"
-    apply (rule has_integral_bound
-       [of _ "\<lambda>x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"])
-    using assms * unfolding has_contour_integral_def
-    apply (auto simp: norm_mult)
-    done
-  then show ?thesis
-    by (auto simp: content_real)
-qed
-
-(*UNUSED
-lemma has_contour_integral_bound_linepath_strong:
-  fixes a :: real and f :: "complex \<Rightarrow> real"
-  assumes "(f has_contour_integral i) (linepath a b)"
-          "finite k"
-          "0 \<le> B" "\<And>x::real. x \<in> closed_segment a b - k \<Longrightarrow> norm(f x) \<le> B"
-    shows "norm i \<le> B*norm(b - a)"
-*)
-
-lemma has_contour_integral_const_linepath: "((\<lambda>x. c) has_contour_integral c*(b - a))(linepath a b)"
-  unfolding has_contour_integral_linepath
-  by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one)
-
-lemma has_contour_integral_0: "((\<lambda>x. 0) has_contour_integral 0) g"
-  by (simp add: has_contour_integral_def)
-
-lemma has_contour_integral_is_0:
-    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> (f has_contour_integral 0) g"
-  by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto
-
-lemma has_contour_integral_setsum:
-    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a has_contour_integral i a) p\<rbrakk>
-     \<Longrightarrow> ((\<lambda>x. setsum (\<lambda>a. f a x) s) has_contour_integral setsum i s) p"
-  by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add)
-
-
-subsection \<open>Operations on path integrals\<close>
-
-lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\<lambda>x. c) = c*(b - a)"
-  by (rule contour_integral_unique [OF has_contour_integral_const_linepath])
-
-lemma contour_integral_neg:
-    "f contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. -(f x)) = -(contour_integral g f)"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg)
-
-lemma contour_integral_add:
-    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x + f2 x) =
-                contour_integral g f1 + contour_integral g f2"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add)
-
-lemma contour_integral_diff:
-    "f1 contour_integrable_on g \<Longrightarrow> f2 contour_integrable_on g \<Longrightarrow> contour_integral g (\<lambda>x. f1 x - f2 x) =
-                contour_integral g f1 - contour_integral g f2"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff)
-
-lemma contour_integral_lmul:
-  shows "f contour_integrable_on g
-           \<Longrightarrow> contour_integral g (\<lambda>x. c * f x) = c*contour_integral g f"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul)
-
-lemma contour_integral_rmul:
-  shows "f contour_integrable_on g
-        \<Longrightarrow> contour_integral g (\<lambda>x. f x * c) = contour_integral g f * c"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul)
-
-lemma contour_integral_div:
-  shows "f contour_integrable_on g
-        \<Longrightarrow> contour_integral g (\<lambda>x. f x / c) = contour_integral g f / c"
-  by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div)
-
-lemma contour_integral_eq:
-    "(\<And>x. x \<in> path_image p \<Longrightarrow> f x = g x) \<Longrightarrow> contour_integral p f = contour_integral p g"
-  apply (simp add: contour_integral_def)
-  using has_contour_integral_eq
-  by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral)
-
-lemma contour_integral_eq_0:
-    "(\<And>z. z \<in> path_image g \<Longrightarrow> f z = 0) \<Longrightarrow> contour_integral g f = 0"
-  by (simp add: has_contour_integral_is_0 contour_integral_unique)
-
-lemma contour_integral_bound_linepath:
-  shows
-    "\<lbrakk>f contour_integrable_on (linepath a b);
-      0 \<le> B; \<And>x. x \<in> closed_segment a b \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
-     \<Longrightarrow> norm(contour_integral (linepath a b) f) \<le> B*norm(b - a)"
-  apply (rule has_contour_integral_bound_linepath [of f])
-  apply (auto simp: has_contour_integral_integral)
-  done
-
-lemma contour_integral_0 [simp]: "contour_integral g (\<lambda>x. 0) = 0"
-  by (simp add: contour_integral_unique has_contour_integral_0)
-
-lemma contour_integral_setsum:
-    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
-     \<Longrightarrow> contour_integral p (\<lambda>x. setsum (\<lambda>a. f a x) s) = setsum (\<lambda>a. contour_integral p (f a)) s"
-  by (auto simp: contour_integral_unique has_contour_integral_setsum has_contour_integral_integral)
-
-lemma contour_integrable_eq:
-    "\<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"
-  unfolding contour_integrable_on_def
-  by (metis has_contour_integral_eq)
-
-
-subsection \<open>Arithmetic theorems for path integrability\<close>
-
-lemma contour_integrable_neg:
-    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. -(f x)) contour_integrable_on g"
-  using has_contour_integral_neg contour_integrable_on_def by blast
-
-lemma contour_integrable_add:
-    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x + f2 x) contour_integrable_on g"
-  using has_contour_integral_add contour_integrable_on_def
-  by fastforce
-
-lemma contour_integrable_diff:
-    "\<lbrakk>f1 contour_integrable_on g; f2 contour_integrable_on g\<rbrakk> \<Longrightarrow> (\<lambda>x. f1 x - f2 x) contour_integrable_on g"
-  using has_contour_integral_diff contour_integrable_on_def
-  by fastforce
-
-lemma contour_integrable_lmul:
-    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. c * f x) contour_integrable_on g"
-  using has_contour_integral_lmul contour_integrable_on_def
-  by fastforce
-
-lemma contour_integrable_rmul:
-    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x * c) contour_integrable_on g"
-  using has_contour_integral_rmul contour_integrable_on_def
-  by fastforce
-
-lemma contour_integrable_div:
-    "f contour_integrable_on g \<Longrightarrow> (\<lambda>x. f x / c) contour_integrable_on g"
-  using has_contour_integral_div contour_integrable_on_def
-  by fastforce
-
-lemma contour_integrable_setsum:
-    "\<lbrakk>finite s; \<And>a. a \<in> s \<Longrightarrow> (f a) contour_integrable_on p\<rbrakk>
-     \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) s) contour_integrable_on p"
-   unfolding contour_integrable_on_def
-   by (metis has_contour_integral_setsum)
-
-
-subsection\<open>Reversing a path integral\<close>
-
-lemma has_contour_integral_reverse_linepath:
-    "(f has_contour_integral i) (linepath a b)
-     \<Longrightarrow> (f has_contour_integral (-i)) (linepath b a)"
-  using has_contour_integral_reversepath valid_path_linepath by fastforce
-
-lemma contour_integral_reverse_linepath:
-    "continuous_on (closed_segment a b) f
-     \<Longrightarrow> contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)"
-apply (rule contour_integral_unique)
-apply (rule has_contour_integral_reverse_linepath)
-by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral)
-
-
-(* Splitting a path integral in a flat way.*)
-
-lemma has_contour_integral_split:
-  assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)"
-      and k: "0 \<le> k" "k \<le> 1"
-      and c: "c - a = k *\<^sub>R (b - a)"
-    shows "(f has_contour_integral (i + j)) (linepath a b)"
-proof (cases "k = 0 \<or> k = 1")
-  case True
-  then show ?thesis
-    using assms
-    apply auto
-    apply (metis add.left_neutral has_contour_integral_trivial has_contour_integral_unique)
-    apply (metis add.right_neutral has_contour_integral_trivial has_contour_integral_unique)
-    done
-next
-  case False
-  then have k: "0 < k" "k < 1" "complex_of_real k \<noteq> 1"
-    using assms apply auto
-    using of_real_eq_iff by fastforce
-  have c': "c = k *\<^sub>R (b - a) + a"
-    by (metis diff_add_cancel c)
-  have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)"
-    by (simp add: algebra_simps c')
-  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}"
-    have **: "\<And>x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b"
-      using False
-      apply (simp add: c' algebra_simps)
-      apply (simp add: real_vector.scale_left_distrib [symmetric] divide_simps)
-      done
-    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}"
-      using * k
-      apply -
-      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse k" "0", simplified])
-      apply (simp_all add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c)
-      apply (drule Integration.has_integral_cmul [where c = "inverse k"])
-      apply (simp add: Integration.has_integral_cmul)
-      done
-  } note fi = this
-  { assume *: "((\<lambda>x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}"
-    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)"
-      using k
-      apply (simp add: c' field_simps)
-      apply (simp add: scaleR_conv_of_real divide_simps)
-      apply (simp add: field_simps)
-      done
-    have "((\<lambda>x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}"
-      using * k
-      apply -
-      apply (drule has_integral_affinity [of _ _ 0 "1::real" "inverse(1 - k)" "-(k/(1 - k))", simplified])
-      apply (simp_all add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc)
-      apply (drule Integration.has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"])
-      apply (simp add: Integration.has_integral_cmul)
-      done
-  } note fj = this
-  show ?thesis
-    using f k
-    apply (simp add: has_contour_integral_linepath)
-    apply (simp add: linepath_def)
-    apply (rule has_integral_combine [OF _ _ fi fj], simp_all)
-    done
-qed
-
-lemma continuous_on_closed_segment_transform:
-  assumes f: "continuous_on (closed_segment a b) f"
-      and k: "0 \<le> k" "k \<le> 1"
-      and c: "c - a = k *\<^sub>R (b - a)"
-    shows "continuous_on (closed_segment a c) f"
-proof -
-  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
-    using c by (simp add: algebra_simps)
-  show "continuous_on (closed_segment a c) f"
-    apply (rule continuous_on_subset [OF f])
-    apply (simp add: segment_convex_hull)
-    apply (rule convex_hull_subset)
-    using assms
-    apply (auto simp: hull_inc c' Convex.convexD_alt)
-    done
-qed
-
-lemma contour_integral_split:
-  assumes f: "continuous_on (closed_segment a b) f"
-      and k: "0 \<le> k" "k \<le> 1"
-      and c: "c - a = k *\<^sub>R (b - a)"
-    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
-proof -
-  have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b"
-    using c by (simp add: algebra_simps)
-  have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f"
-    apply (rule_tac [!] continuous_on_subset [OF f])
-    apply (simp_all add: segment_convex_hull)
-    apply (rule_tac [!] convex_hull_subset)
-    using assms
-    apply (auto simp: hull_inc c' Convex.convexD_alt)
-    done
-  show ?thesis
-    apply (rule contour_integral_unique)
-    apply (rule has_contour_integral_split [OF has_contour_integral_integral has_contour_integral_integral k c])
-    apply (rule contour_integrable_continuous_linepath *)+
-    done
-qed
-
-lemma contour_integral_split_linepath:
-  assumes f: "continuous_on (closed_segment a b) f"
-      and c: "c \<in> closed_segment a b"
-    shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f"
-  using c
-  by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f])
-
-(* The special case of midpoints used in the main quadrisection.*)
-
-lemma has_contour_integral_midpoint:
-  assumes "(f has_contour_integral i) (linepath a (midpoint a b))"
-          "(f has_contour_integral j) (linepath (midpoint a b) b)"
-    shows "(f has_contour_integral (i + j)) (linepath a b)"
-  apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"])
-  using assms
-  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
-  done
-
-lemma contour_integral_midpoint:
-   "continuous_on (closed_segment a b) f
-    \<Longrightarrow> contour_integral (linepath a b) f =
-        contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f"
-  apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"])
-  apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real)
-  done
-
-
-text\<open>A couple of special case lemmas that are useful below\<close>
-
-lemma triangle_linear_has_chain_integral:
-    "((\<lambda>x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-  apply (rule Cauchy_theorem_primitive [of UNIV "\<lambda>x. m/2 * x^2 + d*x"])
-  apply (auto intro!: derivative_eq_intros)
-  done
-
-lemma has_chain_integral_chain_integral3:
-     "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d)
-      \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i"
-  apply (subst contour_integral_unique [symmetric], assumption)
-  apply (drule has_contour_integral_integrable)
-  apply (simp add: valid_path_join)
-  done
-
-lemma has_chain_integral_chain_integral4:
-     "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e)
-      \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i"
-  apply (subst contour_integral_unique [symmetric], assumption)
-  apply (drule has_contour_integral_integrable)
-  apply (simp add: valid_path_join)
-  done
-
-subsection\<open>Reversing the order in a double path integral\<close>
-
-text\<open>The condition is stronger than needed but it's often true in typical situations\<close>
-
-lemma fst_im_cbox [simp]: "cbox c d \<noteq> {} \<Longrightarrow> (fst ` cbox (a,c) (b,d)) = cbox a b"
-  by (auto simp: cbox_Pair_eq)
-
-lemma snd_im_cbox [simp]: "cbox a b \<noteq> {} \<Longrightarrow> (snd ` cbox (a,c) (b,d)) = cbox c d"
-  by (auto simp: cbox_Pair_eq)
-
-lemma contour_integral_swap:
-  assumes fcon:  "continuous_on (path_image g \<times> path_image h) (\<lambda>(y1,y2). f y1 y2)"
-      and vp:    "valid_path g" "valid_path h"
-      and gvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative g (at t))"
-      and hvcon: "continuous_on {0..1} (\<lambda>t. vector_derivative h (at t))"
-  shows "contour_integral g (\<lambda>w. contour_integral h (f w)) =
-         contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
-proof -
-  have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h"
-    using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def)
-  have fgh1: "\<And>x. (\<lambda>t. f (g x) (h t)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g x, h t))"
-    by (rule ext) simp
-  have fgh2: "\<And>x. (\<lambda>t. f (g t) (h x)) = (\<lambda>(y1,y2). f y1 y2) o (\<lambda>t. (g t, h x))"
-    by (rule ext) simp
-  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)"
-    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
-  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)"
-    by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def)
-  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}"
-    apply (rule integrable_continuous_real)
-    apply (rule continuous_on_mult [OF _ gvcon])
-    apply (subst fgh2)
-    apply (rule fcon_im2 gcon continuous_intros | simp)+
-    done
-  have "(\<lambda>z. vector_derivative g (at (fst z))) = (\<lambda>x. vector_derivative g (at x)) o fst"
-    by auto
-  then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\<lambda>x. vector_derivative g (at (fst x)))"
-    apply (rule ssubst)
-    apply (rule continuous_intros | simp add: gvcon)+
-    done
-  have "(\<lambda>z. vector_derivative h (at (snd z))) = (\<lambda>x. vector_derivative h (at x)) o snd"
-    by auto
-  then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\<lambda>x. vector_derivative h (at (snd x)))"
-    apply (rule ssubst)
-    apply (rule continuous_intros | simp add: hvcon)+
-    done
-  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))"
-    by auto
-  then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\<lambda>x. f (g (fst x)) (h (snd x)))"
-    apply (rule ssubst)
-    apply (rule gcon hcon continuous_intros | simp)+
-    apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon])
-    done
-  have "integral {0..1} (\<lambda>x. contour_integral h (f (g x)) * vector_derivative g (at x)) =
-        integral {0..1} (\<lambda>x. contour_integral h (\<lambda>y. f (g x) y * vector_derivative g (at x)))"
-    apply (rule integral_cong [OF contour_integral_rmul [symmetric]])
-    apply (clarsimp simp: contour_integrable_on)
-    apply (rule integrable_continuous_real)
-    apply (rule continuous_on_mult [OF _ hvcon])
-    apply (subst fgh1)
-    apply (rule fcon_im1 hcon continuous_intros | simp)+
-    done
-  also have "... = integral {0..1}
-                     (\<lambda>y. contour_integral g (\<lambda>x. f x (h y) * vector_derivative h (at y)))"
-    apply (simp only: contour_integral_integral)
-    apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified])
-     apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+
-    unfolding integral_mult_left [symmetric]
-    apply (simp only: mult_ac)
-    done
-  also have "... = contour_integral h (\<lambda>z. contour_integral g (\<lambda>w. f w z))"
-    apply (simp add: contour_integral_integral)
-    apply (rule integral_cong)
-    unfolding integral_mult_left [symmetric]
-    apply (simp add: algebra_simps)
-    done
-  finally show ?thesis
-    by (simp add: contour_integral_integral)
-qed
-
-
-subsection\<open>The key quadrisection step\<close>
-
-lemma norm_sum_half:
-  assumes "norm(a + b) >= e"
-    shows "norm a >= e/2 \<or> norm b >= e/2"
-proof -
-  have "e \<le> norm (- a - b)"
-    by (simp add: add.commute assms norm_minus_commute)
-  thus ?thesis
-    using norm_triangle_ineq4 order_trans by fastforce
-qed
-
-lemma norm_sum_lemma:
-  assumes "e \<le> norm (a + b + c + d)"
-    shows "e / 4 \<le> norm a \<or> e / 4 \<le> norm b \<or> e / 4 \<le> norm c \<or> e / 4 \<le> norm d"
-proof -
-  have "e \<le> norm ((a + b) + (c + d))" using assms
-    by (simp add: algebra_simps)
-  then show ?thesis
-    by (auto dest!: norm_sum_half)
-qed
-
-lemma Cauchy_theorem_quadrisection:
-  assumes f: "continuous_on (convex hull {a,b,c}) f"
-      and dist: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
-      and e: "e * K^2 \<le>
-              norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)"
-  shows "\<exists>a' b' c'.
-           a' \<in> convex hull {a,b,c} \<and> b' \<in> convex hull {a,b,c} \<and> c' \<in> convex hull {a,b,c} \<and>
-           dist a' b' \<le> K/2  \<and>  dist b' c' \<le> K/2  \<and>  dist c' a' \<le> K/2  \<and>
-           e * (K/2)^2 \<le> norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)"
-proof -
-  note divide_le_eq_numeral1 [simp del]
-  define a' where "a' = midpoint b c"
-  define b' where "b' = midpoint c a"
-  define c' where "c' = midpoint a b"
-  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
-    using f continuous_on_subset segments_subset_convex_hull by metis+
-  have fcont': "continuous_on (closed_segment c' b') f"
-               "continuous_on (closed_segment a' c') f"
-               "continuous_on (closed_segment b' a') f"
-    unfolding a'_def b'_def c'_def
-    apply (rule continuous_on_subset [OF f],
-           metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+
-    done
-  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
-  have *: "?pathint a b + ?pathint b c + ?pathint c a =
-          (?pathint a c' + ?pathint c' b' + ?pathint b' a) +
-          (?pathint a' c' + ?pathint c' b + ?pathint b a') +
-          (?pathint a' c + ?pathint c b' + ?pathint b' a') +
-          (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
-    apply (simp add: fcont' contour_integral_reverse_linepath)
-    apply (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc)
-    done
-  have [simp]: "\<And>x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2"
-    by (metis left_diff_distrib mult.commute norm_mult_numeral1)
-  have [simp]: "\<And>x y. cmod (x - y) = cmod (y - x)"
-    by (simp add: norm_minus_commute)
-  consider "e * K\<^sup>2 / 4 \<le> cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" |
-           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" |
-           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" |
-           "e * K\<^sup>2 / 4 \<le> cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')"
-    using assms
-    apply (simp only: *)
-    apply (blast intro: that dest!: norm_sum_lemma)
-    done
-  then show ?thesis
-  proof cases
-    case 1 then show ?thesis
-      apply (rule_tac x=a in exI)
-      apply (rule exI [where x=c'])
-      apply (rule exI [where x=b'])
-      using assms
-      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
-      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
-      done
-  next
-    case 2 then show ?thesis
-      apply (rule_tac x=a' in exI)
-      apply (rule exI [where x=c'])
-      apply (rule exI [where x=b])
-      using assms
-      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
-      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
-      done
-  next
-    case 3 then show ?thesis
-      apply (rule_tac x=a' in exI)
-      apply (rule exI [where x=c])
-      apply (rule exI [where x=b'])
-      using assms
-      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
-      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
-      done
-  next
-    case 4 then show ?thesis
-      apply (rule_tac x=a' in exI)
-      apply (rule exI [where x=b'])
-      apply (rule exI [where x=c'])
-      using assms
-      apply (auto simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD])
-      apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real divide_simps)
-      done
-  qed
-qed
-
-subsection\<open>Cauchy's theorem for triangles\<close>
-
-lemma triangle_points_closer:
-  fixes a::complex
-  shows "\<lbrakk>x \<in> convex hull {a,b,c};  y \<in> convex hull {a,b,c}\<rbrakk>
-         \<Longrightarrow> norm(x - y) \<le> norm(a - b) \<or>
-             norm(x - y) \<le> norm(b - c) \<or>
-             norm(x - y) \<le> norm(c - a)"
-  using simplex_extremal_le [of "{a,b,c}"]
-  by (auto simp: norm_minus_commute)
-
-lemma holomorphic_point_small_triangle:
-  assumes x: "x \<in> s"
-      and f: "continuous_on s f"
-      and cd: "f field_differentiable (at x within s)"
-      and e: "0 < e"
-    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>
-              x \<in> convex hull {a,b,c} \<and> convex hull {a,b,c} \<subseteq> s
-              \<longrightarrow> norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f +
-                       contour_integral(linepath c a) f)
-                  \<le> e*(dist a b + dist b c + dist c a)^2"
-           (is "\<exists>k>0. \<forall>a b c. _ \<longrightarrow> ?normle a b c")
-proof -
-  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>
-                     \<Longrightarrow> a \<le> e*(x + y + z)^2"
-    by (simp add: algebra_simps power2_eq_square)
-  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"
-             for x::real and a b c
-    by linarith
-  have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a"
-              if "convex hull {a, b, c} \<subseteq> s" for a b c
-    using segments_subset_convex_hull that
-    by (metis continuous_on_subset f contour_integrable_continuous_linepath)+
-  note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral]
-  { fix f' a b c d
-    assume d: "0 < d"
-       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)"
-       and le: "cmod (a - b) \<le> d" "cmod (b - c) \<le> d" "cmod (c - a) \<le> d"
-       and xc: "x \<in> convex hull {a, b, c}"
-       and s: "convex hull {a, b, c} \<subseteq> s"
-    have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f =
-              contour_integral (linepath a b) (\<lambda>y. f y - f x - f'*(y - x)) +
-              contour_integral (linepath b c) (\<lambda>y. f y - f x - f'*(y - x)) +
-              contour_integral (linepath c a) (\<lambda>y. f y - f x - f'*(y - x))"
-      apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF s])
-      apply (simp add: field_simps)
-      done
-    { fix y
-      assume yc: "y \<in> convex hull {a,b,c}"
-      have "cmod (f y - f x - f' * (y - x)) \<le> e*norm(y - x)"
-        apply (rule f')
-        apply (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans)
-        using s yc by blast
-      also have "... \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))"
-        by (simp add: yc e xc disj_le [OF triangle_points_closer])
-      finally have "cmod (f y - f x - f' * (y - x)) \<le> e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" .
-    } note cm_le = this
-    have "?normle a b c"
-      apply (simp add: dist_norm pa)
-      apply (rule le_of_3)
-      using f' xc s e
-      apply simp_all
-      apply (intro norm_triangle_le add_mono path_bound)
-      apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc)
-      apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+
-      done
-  } note * = this
-  show ?thesis
-    using cd e
-    apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def)
-    apply (clarify dest!: spec mp)
-    using *
-    apply (simp add: dist_norm, blast)
-    done
-qed
-
-
-(* Hence the most basic theorem for a triangle.*)
-locale Chain =
-  fixes x0 At Follows
-  assumes At0: "At x0 0"
-      and AtSuc: "\<And>x n. At x n \<Longrightarrow> \<exists>x'. At x' (Suc n) \<and> Follows x' x"
-begin
-  primrec f where
-    "f 0 = x0"
-  | "f (Suc n) = (SOME x. At x (Suc n) \<and> Follows x (f n))"
-
-  lemma At: "At (f n) n"
-  proof (induct n)
-    case 0 show ?case
-      by (simp add: At0)
-  next
-    case (Suc n) show ?case
-      by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex)
-  qed
-
-  lemma Follows: "Follows (f(Suc n)) (f n)"
-    by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex)
-
-  declare f.simps(2) [simp del]
-end
-
-lemma Chain3:
-  assumes At0: "At x0 y0 z0 0"
-      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"
-  obtains f g h where
-    "f 0 = x0" "g 0 = y0" "h 0 = z0"
-                      "\<And>n. At (f n) (g n) (h n) n"
-                       "\<And>n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)"
-proof -
-  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"
-    apply unfold_locales
-    using At0 AtSuc by auto
-  show ?thesis
-  apply (rule that [of "\<lambda>n. fst (three.f n)"  "\<lambda>n. fst (snd (three.f n))" "\<lambda>n. snd (snd (three.f n))"])
-  apply simp_all
-  using three.At three.Follows
-  apply (simp_all add: split_beta')
-  done
-qed
-
-lemma Cauchy_theorem_triangle:
-  assumes "f holomorphic_on (convex hull {a,b,c})"
-    shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-proof -
-  have contf: "continuous_on (convex hull {a,b,c}) f"
-    by (metis assms holomorphic_on_imp_continuous_on)
-  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
-  { fix y::complex
-    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
-       and ynz: "y \<noteq> 0"
-    define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))"
-    define e where "e = norm y / K^2"
-    have K1: "K \<ge> 1"  by (simp add: K_def max.coboundedI1)
-    then have K: "K > 0" by linarith
-    have [iff]: "dist a b \<le> K" "dist b c \<le> K" "dist c a \<le> K"
-      by (simp_all add: K_def)
-    have e: "e > 0"
-      unfolding e_def using ynz K1 by simp
-    define At where "At x y z n \<longleftrightarrow>
-        convex hull {x,y,z} \<subseteq> convex hull {a,b,c} \<and>
-        dist x y \<le> K/2^n \<and> dist y z \<le> K/2^n \<and> dist z x \<le> K/2^n \<and>
-        norm(?pathint x y + ?pathint y z + ?pathint z x) \<ge> e*(K/2^n)^2"
-      for x y z n
-    have At0: "At a b c 0"
-      using fy
-      by (simp add: At_def e_def has_chain_integral_chain_integral3)
-    { fix x y z n
-      assume At: "At x y z n"
-      then have contf': "continuous_on (convex hull {x,y,z}) f"
-        using contf At_def continuous_on_subset by blast
-      have "\<exists>x' y' z'. At x' y' z' (Suc n) \<and> convex hull {x',y',z'} \<subseteq> convex hull {x,y,z}"
-        using At
-        apply (simp add: At_def)
-        using  Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e]
-        apply clarsimp
-        apply (rule_tac x="a'" in exI)
-        apply (rule_tac x="b'" in exI)
-        apply (rule_tac x="c'" in exI)
-        apply (simp add: algebra_simps)
-        apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE)
-        done
-    } note AtSuc = this
-    obtain fa fb fc
-      where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c"
-        and cosb: "\<And>n. convex hull {fa n, fb n, fc n} \<subseteq> convex hull {a,b,c}"
-        and dist: "\<And>n. dist (fa n) (fb n) \<le> K/2^n"
-                  "\<And>n. dist (fb n) (fc n) \<le> K/2^n"
-                  "\<And>n. dist (fc n) (fa n) \<le> K/2^n"
-        and no: "\<And>n. norm(?pathint (fa n) (fb n) +
-                           ?pathint (fb n) (fc n) +
-                           ?pathint (fc n) (fa n)) \<ge> e * (K/2^n)^2"
-        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}"
-      apply (rule Chain3 [of At, OF At0 AtSuc])
-      apply (auto simp: At_def)
-      done
-    have "\<exists>x. \<forall>n. x \<in> convex hull {fa n, fb n, fc n}"
-      apply (rule bounded_closed_nest)
-      apply (simp_all add: compact_imp_closed finite_imp_compact_convex_hull finite_imp_bounded_convex_hull)
-      apply (rule allI)
-      apply (rule transitive_stepwise_le)
-      apply (auto simp: conv_le)
-      done
-    then obtain x where x: "\<And>n. x \<in> convex hull {fa n, fb n, fc n}" by auto
-    then have xin: "x \<in> convex hull {a,b,c}"
-      using assms f0 by blast
-    then have fx: "f field_differentiable at x within (convex hull {a,b,c})"
-      using assms holomorphic_on_def by blast
-    { fix k n
-      assume k: "0 < k"
-         and le:
-            "\<And>x' y' z'.
-               \<lbrakk>dist x' y' \<le> k; dist y' z' \<le> k; dist z' x' \<le> k;
-                x \<in> convex hull {x',y',z'};
-                convex hull {x',y',z'} \<subseteq> convex hull {a,b,c}\<rbrakk>
-               \<Longrightarrow>
-               cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10
-                     \<le> e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2"
-         and Kk: "K / k < 2 ^ n"
-      have "K / 2 ^ n < k" using Kk k
-        by (auto simp: field_simps)
-      then have DD: "dist (fa n) (fb n) \<le> k" "dist (fb n) (fc n) \<le> k" "dist (fc n) (fa n) \<le> k"
-        using dist [of n]  k
-        by linarith+
-      have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2
-               \<le> (3 * K / 2 ^ n)\<^sup>2"
-        using dist [of n] e K
-        by (simp add: abs_le_square_iff [symmetric])
-      have less10: "\<And>x y::real. 0 < x \<Longrightarrow> y \<le> 9*x \<Longrightarrow> y < x*10"
-        by linarith
-      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"
-        using ynz dle e mult_le_cancel_left_pos by blast
-      also have "... <
-          cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10"
-        using no [of n] e K
-        apply (simp add: e_def field_simps)
-        apply (simp only: zero_less_norm_iff [symmetric])
-        done
-      finally have False
-        using le [OF DD x cosb] by auto
-    } then
-    have ?thesis
-      using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e
-      apply clarsimp
-      apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]])
-      apply force+
-      done
-  }
-  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
-    by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1)
-                   segments_subset_convex_hull(3) segments_subset_convex_hull(5))
-  ultimately show ?thesis
-    using has_contour_integral_integral by fastforce
-qed
-
-
-subsection\<open>Version needing function holomorphic in interior only\<close>
-
-lemma Cauchy_theorem_flat_lemma:
-  assumes f: "continuous_on (convex hull {a,b,c}) f"
-      and c: "c - a = k *\<^sub>R (b - a)"
-      and k: "0 \<le> k"
-    shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f +
-          contour_integral (linepath c a) f = 0"
-proof -
-  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
-    using f continuous_on_subset segments_subset_convex_hull by metis+
-  show ?thesis
-  proof (cases "k \<le> 1")
-    case True show ?thesis
-      by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc)
-  next
-    case False then show ?thesis
-      using fabc c
-      apply (subst contour_integral_split [of a c f "1/k" b, symmetric])
-      apply (metis closed_segment_commute fabc(3))
-      apply (auto simp: k contour_integral_reverse_linepath)
-      done
-  qed
-qed
-
-lemma Cauchy_theorem_flat:
-  assumes f: "continuous_on (convex hull {a,b,c}) f"
-      and c: "c - a = k *\<^sub>R (b - a)"
-    shows "contour_integral (linepath a b) f +
-           contour_integral (linepath b c) f +
-           contour_integral (linepath c a) f = 0"
-proof (cases "0 \<le> k")
-  case True with assms show ?thesis
-    by (blast intro: Cauchy_theorem_flat_lemma)
-next
-  case False
-  have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
-    using f continuous_on_subset segments_subset_convex_hull by metis+
-  moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f +
-        contour_integral (linepath c b) f = 0"
-    apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"])
-    using False c
-    apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps)
-    done
-  ultimately show ?thesis
-    apply (auto simp: contour_integral_reverse_linepath)
-    using add_eq_0_iff by force
-qed
-
-
-lemma Cauchy_theorem_triangle_interior:
-  assumes contf: "continuous_on (convex hull {a,b,c}) f"
-      and holf:  "f holomorphic_on interior (convex hull {a,b,c})"
-     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-proof -
-  have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f"
-    using contf continuous_on_subset segments_subset_convex_hull by metis+
-  have "bounded (f ` (convex hull {a,b,c}))"
-    by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf)
-  then obtain B where "0 < B" and Bnf: "\<And>x. x \<in> convex hull {a,b,c} \<Longrightarrow> norm (f x) \<le> B"
-     by (auto simp: dest!: bounded_pos [THEN iffD1])
-  have "bounded (convex hull {a,b,c})"
-    by (simp add: bounded_convex_hull)
-  then obtain C where C: "0 < C" and Cno: "\<And>y. y \<in> convex hull {a,b,c} \<Longrightarrow> norm y < C"
-    using bounded_pos_less by blast
-  then have diff_2C: "norm(x - y) \<le> 2*C"
-           if x: "x \<in> convex hull {a, b, c}" and y: "y \<in> convex hull {a, b, c}" for x y
-  proof -
-    have "cmod x \<le> C"
-      using x by (meson Cno not_le not_less_iff_gr_or_eq)
-    hence "cmod (x - y) \<le> C + C"
-      using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans)
-    thus "cmod (x - y) \<le> 2 * C"
-      by (metis mult_2)
-  qed
-  have contf': "continuous_on (convex hull {b,a,c}) f"
-    using contf by (simp add: insert_commute)
-  { fix y::complex
-    assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
-       and ynz: "y \<noteq> 0"
-    have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y"
-      by (rule has_chain_integral_chain_integral3 [OF fy])
-    have ?thesis
-    proof (cases "c=a \<or> a=b \<or> b=c")
-      case True then show ?thesis
-        using Cauchy_theorem_flat [OF contf, of 0]
-        using has_chain_integral_chain_integral3 [OF fy] ynz
-        by (force simp: fabc contour_integral_reverse_linepath)
-    next
-      case False
-      then have car3: "card {a, b, c} = Suc (DIM(complex))"
-        by auto
-      { assume "interior(convex hull {a,b,c}) = {}"
-        then have "collinear{a,b,c}"
-          using interior_convex_hull_eq_empty [OF car3]
-          by (simp add: collinear_3_eq_affine_dependent)
-        then have "False"
-          using False
-          apply (clarsimp simp add: collinear_3 collinear_lemma)
-          apply (drule Cauchy_theorem_flat [OF contf'])
-          using pi_eq_y ynz
-          apply (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath)
-          done
-      }
-      then obtain d where d: "d \<in> interior (convex hull {a, b, c})"
-        by blast
-      { fix d1
-        assume d1_pos: "0 < d1"
-           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>
-                           \<Longrightarrow> cmod (f x' - f x) < cmod y / (24 * C)"
-        define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))"
-        define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x
-        let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
-        have e: "0 < e" "e \<le> 1" "e \<le> d1 / (4 * C)" "e \<le> cmod y / 24 / C / B"
-          using d1_pos \<open>C>0\<close> \<open>B>0\<close> ynz by (simp_all add: e_def)
-        then have eCB: "24 * e * C * B \<le> cmod y"
-          using \<open>C>0\<close> \<open>B>0\<close>  by (simp add: field_simps)
-        have e_le_d1: "e * (4 * C) \<le> d1"
-          using e \<open>C>0\<close> by (simp add: field_simps)
-        have "shrink a \<in> interior(convex hull {a,b,c})"
-             "shrink b \<in> interior(convex hull {a,b,c})"
-             "shrink c \<in> interior(convex hull {a,b,c})"
-          using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
-        then have fhp0: "(f has_contour_integral 0)
-                (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))"
-          by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal convex_interior)
-        then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0"
-          by (simp add: has_chain_integral_chain_integral3)
-        have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)"
-                      "f contour_integrable_on linepath (shrink b) (shrink c)"
-                      "f contour_integrable_on linepath (shrink c) (shrink a)"
-          using fhp0  by (auto simp: valid_path_join dest: has_contour_integral_integrable)
-        have cmod_shr: "\<And>x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)"
-          using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric])
-        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)"
-          by (simp add: algebra_simps)
-        have "cmod y / (24 * C) \<le> cmod y / cmod (b - a) / 12"
-          using False \<open>C>0\<close> diff_2C [of b a] ynz
-          by (auto simp: divide_simps hull_inc)
-        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
-          apply (cases "x=0", simp add: \<open>0<C\<close>)
-          using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast
-        { fix u v
-          assume uv: "u \<in> convex hull {a, b, c}" "v \<in> convex hull {a, b, c}" "u\<noteq>v"
-             and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)"
-          have shr_uv: "shrink u \<in> interior(convex hull {a,b,c})"
-                       "shrink v \<in> interior(convex hull {a,b,c})"
-            using d e uv
-            by (auto simp: hull_inc mem_interior_convex_shrink shrink_def)
-          have cmod_fuv: "\<And>x. 0\<le>x \<Longrightarrow> x\<le>1 \<Longrightarrow> cmod (f (linepath (shrink u) (shrink v) x)) \<le> B"
-            using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD])
-          have By_uv: "B * (12 * (e * cmod (u - v))) \<le> cmod y"
-            apply (rule order_trans [OF _ eCB])
-            using e \<open>B>0\<close> diff_2C [of u v] uv
-            by (auto simp: field_simps)
-          { fix x::real   assume x: "0\<le>x" "x\<le>1"
-            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)"
-              apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0])
-              using uv x d interior_subset
-              apply (auto simp: hull_inc intro!: less_C)
-              done
-            have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))"
-              by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real)
-            have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1"
-              using \<open>e>0\<close>
-              apply (simp add: ll norm_mult scaleR_diff_right)
-              apply (rule less_le_trans [OF _ e_le_d1])
-              using cmod_less_4C
-              apply (force intro: norm_triangle_lt)
-              done
-            have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)"
-              using x uv shr_uv cmod_less_dt
-              by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull)
-            also have "... \<le> cmod y / cmod (v - u) / 12"
-              using False uv \<open>C>0\<close> diff_2C [of v u] ynz
-              by (auto simp: divide_simps hull_inc)
-            finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \<le> cmod y / cmod (v - u) / 12"
-              by simp
-            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"
-              using uv False by (auto simp: field_simps)
-            have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) +
-                  cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x))
-                  \<le> cmod y / 6"
-              apply (rule order_trans [of _ "B*((norm y / 24 / C / B)*2*C) + (2*C)*(norm y /24 / C)"])
-              apply (rule add_mono [OF mult_mono])
-              using By_uv e \<open>0 < B\<close> \<open>0 < C\<close> x ynz
-              apply (simp_all add: cmod_fuv cmod_shr cmod_12_le hull_inc)
-              apply (simp add: field_simps)
-              done
-          } note cmod_diff_le = this
-          have f_uv: "continuous_on (closed_segment u v) f"
-            by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull])
-          have **: "\<And>f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)"
-            by (simp add: algebra_simps)
-          have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \<le> norm y / 6"
-            apply (rule order_trans)
-            apply (rule has_integral_bound
-                    [of "B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)"
-                        "\<lambda>x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)"
-                        _ 0 1 ])
-            using ynz \<open>0 < B\<close> \<open>0 < C\<close>
-            apply (simp_all del: le_divide_eq_numeral1)
-            apply (simp add: has_integral_sub has_contour_integral_linepath [symmetric] has_contour_integral_integral
-                             fpi_uv f_uv contour_integrable_continuous_linepath, clarify)
-            apply (simp only: **)
-            apply (simp add: norm_triangle_le norm_mult cmod_diff_le del: le_divide_eq_numeral1)
-            done
-          } note * = this
-          have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \<le> norm y / 6"
-            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
-          moreover
-          have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \<le> norm y / 6"
-            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
-          moreover
-          have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \<le> norm y / 6"
-            using False fpi_abc by (rule_tac *) (auto simp: hull_inc)
-          ultimately
-          have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) +
-                     (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a))
-                \<le> norm y / 6 + norm y / 6 + norm y / 6"
-            by (metis norm_triangle_le add_mono)
-          also have "... = norm y / 2"
-            by simp
-          finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) -
-                          (?pathint a b + ?pathint b c + ?pathint c a))
-                \<le> norm y / 2"
-            by (simp add: algebra_simps)
-          then
-          have "norm(?pathint a b + ?pathint b c + ?pathint c a) \<le> norm y / 2"
-            by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff)
-          then have "False"
-            using pi_eq_y ynz by auto
-        }
-        moreover have "uniformly_continuous_on (convex hull {a,b,c}) f"
-          by (simp add: contf compact_convex_hull compact_uniformly_continuous)
-        ultimately have "False"
-          unfolding uniformly_continuous_on_def
-          by (force simp: ynz \<open>0 < C\<close> dist_norm)
-        then show ?thesis ..
-      qed
-  }
-  moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
-    using fabc contour_integrable_continuous_linepath by auto
-  ultimately show ?thesis
-    using has_contour_integral_integral by fastforce
-qed
-
-
-subsection\<open>Version allowing finite number of exceptional points\<close>
-
-lemma Cauchy_theorem_triangle_cofinite:
-  assumes "continuous_on (convex hull {a,b,c}) f"
-      and "finite s"
-      and "(\<And>x. x \<in> interior(convex hull {a,b,c}) - s \<Longrightarrow> f field_differentiable (at x))"
-     shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-using assms
-proof (induction "card s" arbitrary: a b c s rule: less_induct)
-  case (less s a b c)
-  show ?case
-  proof (cases "s={}")
-    case True with less show ?thesis
-      by (fastforce simp: holomorphic_on_def field_differentiable_at_within
-                    Cauchy_theorem_triangle_interior)
-  next
-    case False
-    then obtain d s' where d: "s = insert d s'" "d \<notin> s'"
-      by (meson Set.set_insert all_not_in_conv)
-    then show ?thesis
-    proof (cases "d \<in> convex hull {a,b,c}")
-      case False
-      show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-        apply (rule less.hyps [of "s'"])
-        using False d \<open>finite s\<close> interior_subset
-        apply (auto intro!: less.prems)
-        done
-    next
-      case True
-      have *: "convex hull {a, b, d} \<subseteq> convex hull {a, b, c}"
-        by (meson True hull_subset insert_subset convex_hull_subset)
-      have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)"
-        apply (rule less.hyps [of "s'"])
-        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
-        apply (auto intro!: less.prems continuous_on_subset [OF  _ *])
-        apply (metis * insert_absorb insert_subset interior_mono)
-        done
-      have *: "convex hull {b, c, d} \<subseteq> convex hull {a, b, c}"
-        by (meson True hull_subset insert_subset convex_hull_subset)
-      have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)"
-        apply (rule less.hyps [of "s'"])
-        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
-        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
-        apply (metis * insert_absorb insert_subset interior_mono)
-        done
-      have *: "convex hull {c, a, d} \<subseteq> convex hull {a, b, c}"
-        by (meson True hull_subset insert_subset convex_hull_subset)
-      have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)"
-        apply (rule less.hyps [of "s'"])
-        using True d  \<open>finite s\<close> not_in_interior_convex_hull_3
-        apply (auto intro!: less.prems continuous_on_subset [OF _ *])
-        apply (metis * insert_absorb insert_subset interior_mono)
-        done
-      have "f contour_integrable_on linepath a b"
-        using less.prems
-        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
-      moreover have "f contour_integrable_on linepath b c"
-        using less.prems
-        by (metis continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(3))
-      moreover have "f contour_integrable_on linepath c a"
-        using less.prems
-        by (metis continuous_on_subset insert_commute contour_integrable_continuous_linepath segments_subset_convex_hull(3))
-      ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)"
-        by auto
-      { fix y::complex
-        assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)"
-           and ynz: "y \<noteq> 0"
-        have cont_ad: "continuous_on (closed_segment a d) f"
-          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3))
-        have cont_bd: "continuous_on (closed_segment b d) f"
-          by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1))
-        have cont_cd: "continuous_on (closed_segment c d) f"
-          by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2))
-        have "contour_integral  (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))"
-                "contour_integral  (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))"
-                "contour_integral  (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)"
-            using has_chain_integral_chain_integral3 [OF abd]
-                  has_chain_integral_chain_integral3 [OF bcd]
-                  has_chain_integral_chain_integral3 [OF cad]
-            by (simp_all add: algebra_simps add_eq_0_iff)
-        then have ?thesis
-          using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce
-      }
-      then show ?thesis
-        using fpi contour_integrable_on_def by blast
-    qed
-  qed
-qed
-
-
-subsection\<open>Cauchy's theorem for an open starlike set\<close>
-
-lemma starlike_convex_subset:
-  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"
-    shows "convex hull {a,b,c} \<subseteq> s"
-      using s
-      apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull)
-      apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE)
-      done
-
-lemma triangle_contour_integrals_starlike_primitive:
-  assumes contf: "continuous_on s f"
-      and s: "a \<in> s" "open s"
-      and x: "x \<in> s"
-      and subs: "\<And>y. y \<in> s \<Longrightarrow> closed_segment a y \<subseteq> s"
-      and zer: "\<And>b c. closed_segment b c \<subseteq> s
-                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
-                       contour_integral (linepath c a) f = 0"
-    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)"
-proof -
-  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
-  { fix e y
-    assume e: "0 < e" and bxe: "ball x e \<subseteq> s" and close: "cmod (y - x) < e"
-    have y: "y \<in> s"
-      using bxe close  by (force simp: dist_norm norm_minus_commute)
-    have cont_ayf: "continuous_on (closed_segment a y) f"
-      using contf continuous_on_subset subs y by blast
-    have xys: "closed_segment x y \<subseteq> s"
-      apply (rule order_trans [OF _ bxe])
-      using close
-      by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound)
-    have "?pathint a y - ?pathint a x = ?pathint x y"
-      using zer [OF xys]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
-  } note [simp] = this
-  { fix e::real
-    assume e: "0 < e"
-    have cont_atx: "continuous (at x) f"
-      using x s contf continuous_on_eq_continuous_at by blast
-    then obtain d1 where d1: "d1>0" and d1_less: "\<And>y. cmod (y - x) < d1 \<Longrightarrow> cmod (f y - f x) < e/2"
-      unfolding continuous_at Lim_at dist_norm  using e
-      by (drule_tac x="e/2" in spec) force
-    obtain d2 where d2: "d2>0" "ball x d2 \<subseteq> s" using  \<open>open s\<close> x
-      by (auto simp: open_contains_ball)
-    have dpos: "min d1 d2 > 0" using d1 d2 by simp
-    { fix y
-      assume yx: "y \<noteq> x" and close: "cmod (y - x) < min d1 d2"
-      have y: "y \<in> s"
-        using d2 close  by (force simp: dist_norm norm_minus_commute)
-      have fxy: "f contour_integrable_on linepath x y"
-        apply (rule contour_integrable_continuous_linepath)
-        apply (rule continuous_on_subset [OF contf])
-        using close d2
-        apply (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1))
-        done
-      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
-        by (auto simp: contour_integrable_on_def)
-      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
-        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
-      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
-        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
-        using e apply simp
-        apply (rule d1_less [THEN less_imp_le])
-        using close segment_bound
-        apply force
-        done
-      also have "... < e * cmod (y - x)"
-        by (simp add: e yx)
-      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
-        using i yx  by (simp add: contour_integral_unique divide_less_eq)
-    }
-    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"
-      using dpos by blast
-  }
-  then have *: "(\<lambda>y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \<midarrow>x\<rightarrow> 0"
-    by (simp add: Lim_at dist_norm inverse_eq_divide)
-  show ?thesis
-    apply (simp add: has_field_derivative_def has_derivative_at bounded_linear_mult_right)
-    apply (rule Lim_transform [OF * Lim_eventually])
-    apply (simp add: inverse_eq_divide [symmetric] eventually_at)
-    using \<open>open s\<close> x
-    apply (force simp: dist_norm open_contains_ball)
-    done
-qed
-
-(** Existence of a primitive.*)
-
-lemma holomorphic_starlike_primitive:
-  fixes f :: "complex \<Rightarrow> complex"
-  assumes contf: "continuous_on s f"
-      and s: "starlike s" and os: "open s"
-      and k: "finite k"
-      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
-    shows "\<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x)"
-proof -
-  obtain a where a: "a\<in>s" and a_cs: "\<And>x. x\<in>s \<Longrightarrow> closed_segment a x \<subseteq> s"
-    using s by (auto simp: starlike_def)
-  { fix x b c
-    assume "x \<in> s" "closed_segment b c \<subseteq> s"
-    then have abcs: "convex hull {a, b, c} \<subseteq> s"
-      by (simp add: a a_cs starlike_convex_subset)
-    then have *: "continuous_on (convex hull {a, b, c}) f"
-      by (simp add: continuous_on_subset [OF contf])
-    have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)"
-      apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
-      using abcs apply (simp add: continuous_on_subset [OF contf])
-      using * abcs interior_subset apply (auto intro: fcd)
-      done
-  } note 0 = this
-  show ?thesis
-    apply (intro exI ballI)
-    apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption)
-    apply (metis a_cs)
-    apply (metis has_chain_integral_chain_integral3 0)
-    done
-qed
-
-lemma Cauchy_theorem_starlike:
- "\<lbrakk>open s; starlike s; finite k; continuous_on s f;
-   \<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x;
-   valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
-   \<Longrightarrow> (f has_contour_integral 0)  g"
-  by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open)
-
-lemma Cauchy_theorem_starlike_simple:
-  "\<lbrakk>open s; starlike s; f holomorphic_on s; valid_path g; path_image g \<subseteq> s; pathfinish g = pathstart g\<rbrakk>
-   \<Longrightarrow> (f has_contour_integral 0) g"
-apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI])
-apply (simp_all add: holomorphic_on_imp_continuous_on)
-apply (metis at_within_open holomorphic_on_def)
-done
-
-
-subsection\<open>Cauchy's theorem for a convex set\<close>
-
-text\<open>For a convex set we can avoid assuming openness and boundary analyticity\<close>
-
-lemma triangle_contour_integrals_convex_primitive:
-  assumes contf: "continuous_on s f"
-      and s: "a \<in> s" "convex s"
-      and x: "x \<in> s"
-      and zer: "\<And>b c. \<lbrakk>b \<in> s; c \<in> s\<rbrakk>
-                   \<Longrightarrow> contour_integral (linepath a b) f + contour_integral (linepath b c) f +
-                       contour_integral (linepath c a) f = 0"
-    shows "((\<lambda>x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within s)"
-proof -
-  let ?pathint = "\<lambda>x y. contour_integral(linepath x y) f"
-  { fix y
-    assume y: "y \<in> s"
-    have cont_ayf: "continuous_on (closed_segment a y) f"
-      using s y  by (meson contf continuous_on_subset convex_contains_segment)
-    have xys: "closed_segment x y \<subseteq> s"  (*?*)
-      using convex_contains_segment s x y by auto
-    have "?pathint a y - ?pathint a x = ?pathint x y"
-      using zer [OF x y]  contour_integral_reverse_linepath [OF cont_ayf]  add_eq_0_iff by force
-  } note [simp] = this
-  { fix e::real
-    assume e: "0 < e"
-    have cont_atx: "continuous (at x within s) f"
-      using x s contf  by (simp add: continuous_on_eq_continuous_within)
-    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"
-      unfolding continuous_within Lim_within dist_norm using e
-      by (drule_tac x="e/2" in spec) force
-    { fix y
-      assume yx: "y \<noteq> x" and close: "cmod (y - x) < d1" and y: "y \<in> s"
-      have fxy: "f contour_integrable_on linepath x y"
-        using convex_contains_segment s x y
-        by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
-      then obtain i where i: "(f has_contour_integral i) (linepath x y)"
-        by (auto simp: contour_integrable_on_def)
-      then have "((\<lambda>w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)"
-        by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath])
-      then have "cmod (i - f x * (y - x)) \<le> e / 2 * cmod (y - x)"
-        apply (rule has_contour_integral_bound_linepath [where B = "e/2"])
-        using e apply simp
-        apply (rule d1_less [THEN less_imp_le])
-        using convex_contains_segment s(2) x y apply blast
-        using close segment_bound(1) apply fastforce
-        done
-      also have "... < e * cmod (y - x)"
-        by (simp add: e yx)
-      finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e"
-        using i yx  by (simp add: contour_integral_unique divide_less_eq)
-    }
-    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"
-      using d1 by blast
-  }
-  then have *: "((\<lambda>y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \<longlongrightarrow> 0) (at x within s)"
-    by (simp add: Lim_within dist_norm inverse_eq_divide)
-  show ?thesis
-    apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right)
-    apply (rule Lim_transform [OF * Lim_eventually])
-    using linordered_field_no_ub
-    apply (force simp: inverse_eq_divide [symmetric] eventually_at)
-    done
-qed
-
-lemma contour_integral_convex_primitive:
-  "\<lbrakk>convex s; continuous_on s f;
-    \<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>
-         \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
-  apply (cases "s={}")
-  apply (simp_all add: ex_in_conv [symmetric])
-  apply (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3)
-  done
-
-lemma holomorphic_convex_primitive:
-  fixes f :: "complex \<Rightarrow> complex"
-  shows
-  "\<lbrakk>convex s; finite k; continuous_on s f;
-    \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x\<rbrakk>
-   \<Longrightarrow> \<exists>g. \<forall>x \<in> s. (g has_field_derivative f x) (at x within s)"
-apply (rule contour_integral_convex_primitive [OF _ _ Cauchy_theorem_triangle_cofinite])
-prefer 3
-apply (erule continuous_on_subset)
-apply (simp add: subset_hull continuous_on_subset, assumption+)
-by (metis Diff_iff convex_contains_segment insert_absorb insert_subset interior_mono segment_convex_hull subset_hull)
-
-lemma Cauchy_theorem_convex:
-    "\<lbrakk>continuous_on s f; convex s; finite k;
-      \<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x;
-     valid_path g; path_image g \<subseteq> s;
-     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
-  by (metis holomorphic_convex_primitive Cauchy_theorem_primitive)
-
-lemma Cauchy_theorem_convex_simple:
-    "\<lbrakk>f holomorphic_on s; convex s;
-     valid_path g; path_image g \<subseteq> s;
-     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
-  apply (rule Cauchy_theorem_convex)
-  apply (simp_all add: holomorphic_on_imp_continuous_on)
-  apply (rule finite.emptyI)
-  using at_within_interior holomorphic_on_def interior_subset by fastforce
-
-
-text\<open>In particular for a disc\<close>
-lemma Cauchy_theorem_disc:
-    "\<lbrakk>finite k; continuous_on (cball a e) f;
-      \<And>x. x \<in> ball a e - k \<Longrightarrow> f field_differentiable at x;
-     valid_path g; path_image g \<subseteq> cball a e;
-     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
-  apply (rule Cauchy_theorem_convex)
-  apply (auto simp: convex_cball interior_cball)
-  done
-
-lemma Cauchy_theorem_disc_simple:
-    "\<lbrakk>f holomorphic_on (ball a e); valid_path g; path_image g \<subseteq> ball a e;
-     pathfinish g = pathstart g\<rbrakk> \<Longrightarrow> (f has_contour_integral 0) g"
-by (simp add: Cauchy_theorem_convex_simple)
-
-
-subsection\<open>Generalize integrability to local primitives\<close>
-
-lemma contour_integral_local_primitive_lemma:
-  fixes f :: "complex\<Rightarrow>complex"
-  shows
-    "\<lbrakk>g piecewise_differentiable_on {a..b};
-      \<And>x. x \<in> s \<Longrightarrow> (f has_field_derivative f' x) (at x within s);
-      \<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s\<rbrakk>
-     \<Longrightarrow> (\<lambda>x. f' (g x) * vector_derivative g (at x within {a..b}))
-            integrable_on {a..b}"
-  apply (cases "cbox a b = {}", force)
-  apply (simp add: integrable_on_def)
-  apply (rule exI)
-  apply (rule contour_integral_primitive_lemma, assumption+)
-  using atLeastAtMost_iff by blast
-
-lemma contour_integral_local_primitive_any:
-  fixes f :: "complex \<Rightarrow> complex"
-  assumes gpd: "g piecewise_differentiable_on {a..b}"
-      and dh: "\<And>x. x \<in> s
-               \<Longrightarrow> \<exists>d h. 0 < d \<and>
-                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
-      and gs: "\<And>x. x \<in> {a..b} \<Longrightarrow> g x \<in> s"
-  shows "(\<lambda>x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}"
-proof -
-  { fix x
-    assume x: "a \<le> x" "x \<le> b"
-    obtain d h where d: "0 < d"
-               and h: "(\<And>y. norm(y - g x) < d \<Longrightarrow> (h has_field_derivative f y) (at y within s))"
-      using x gs dh by (metis atLeastAtMost_iff)
-    have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast
-    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"
-      using x d
-      apply (auto simp: dist_norm continuous_on_iff)
-      apply (drule_tac x=x in bspec)
-      using x apply simp
-      apply (drule_tac x=d in spec, auto)
-      done
-    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>
-                          (\<lambda>x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}"
-      apply (rule_tac x=e in exI)
-      using e
-      apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify)
-      apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma)
-        apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset)
-       apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force)
-      done
-  } then
-  show ?thesis
-    by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified])
-qed
-
-lemma contour_integral_local_primitive:
-  fixes f :: "complex \<Rightarrow> complex"
-  assumes g: "valid_path g" "path_image g \<subseteq> s"
-      and dh: "\<And>x. x \<in> s
-               \<Longrightarrow> \<exists>d h. 0 < d \<and>
-                         (\<forall>y. norm(y - x) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
-  shows "f contour_integrable_on g"
-  using g
-  apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def
-            has_integral_localized_vector_derivative integrable_on_def [symmetric])
-  using contour_integral_local_primitive_any [OF _ dh]
-  by (meson image_subset_iff piecewise_C1_imp_differentiable)
-
-
-text\<open>In particular if a function is holomorphic\<close>
-
-lemma contour_integrable_holomorphic:
-  assumes contf: "continuous_on s f"
-      and os: "open s"
-      and k: "finite k"
-      and g: "valid_path g" "path_image g \<subseteq> s"
-      and fcd: "\<And>x. x \<in> s - k \<Longrightarrow> f field_differentiable at x"
-    shows "f contour_integrable_on g"
-proof -
-  { fix z
-    assume z: "z \<in> s"
-    obtain d where d: "d>0" "ball z d \<subseteq> s" using  \<open>open s\<close> z
-      by (auto simp: open_contains_ball)
-    then have contfb: "continuous_on (ball z d) f"
-      using contf continuous_on_subset by blast
-    obtain h where "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within ball z d)"
-      using holomorphic_convex_primitive [OF convex_ball k contfb fcd] d
-            interior_subset by force
-    then have "\<forall>y\<in>ball z d. (h has_field_derivative f y) (at y within s)"
-      by (metis Topology_Euclidean_Space.open_ball at_within_open d(2) os subsetCE)
-    then have "\<exists>h. (\<forall>y. cmod (y - z) < d \<longrightarrow> (h has_field_derivative f y) (at y within s))"
-      by (force simp: dist_norm norm_minus_commute)
-    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))"
-      using d by blast
-  }
-  then show ?thesis
-    by (rule contour_integral_local_primitive [OF g])
-qed
-
-lemma contour_integrable_holomorphic_simple:
-  assumes fh: "f holomorphic_on s"
-      and os: "open s"
-      and g: "valid_path g" "path_image g \<subseteq> s"
-    shows "f contour_integrable_on g"
-  apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g])
-  apply (simp add: fh holomorphic_on_imp_continuous_on)
-  using fh  by (simp add: field_differentiable_def holomorphic_on_open os)
-
-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 contour_integrable_inversediff:
-    "\<lbrakk>valid_path g; z \<notin> path_image g\<rbrakk> \<Longrightarrow> (\<lambda>w. 1 / (w-z)) contour_integrable_on g"
-apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"])
-apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros)
-done
-
-text\<open>Key fact that path integral is the same for a "nearby" path. This is the
- main lemma for the homotopy form of Cauchy's theorem and is also useful
- if we want "without loss of generality" to assume some nice properties of a
- path (e.g. smoothness). It can also be used to define the integrals of
- analytic functions over arbitrary continuous paths. This is just done for
- winding numbers now.
-\<close>
-
-text\<open>A technical definition to avoid duplication of similar proofs,
-     for paths joined at the ends versus looping paths\<close>
-definition linked_paths :: "bool \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
-  where "linked_paths atends g h ==
-        (if atends then pathstart h = pathstart g \<and> pathfinish h = pathfinish g
-                   else pathfinish g = pathstart g \<and> pathfinish h = pathstart h)"
-
-text\<open>This formulation covers two cases: @{term g} and @{term h} share their
-      start and end points; @{term g} and @{term h} both loop upon themselves.\<close>
-lemma contour_integral_nearby:
-  assumes os: "open s" and p: "path p" "path_image p \<subseteq> s"
-    shows
-       "\<exists>d. 0 < d \<and>
-            (\<forall>g h. valid_path g \<and> valid_path h \<and>
-                  (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
-                  linked_paths atends g h
-                  \<longrightarrow> 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))"
-proof -
-  have "\<forall>z. \<exists>e. z \<in> path_image p \<longrightarrow> 0 < e \<and> ball z e \<subseteq> s"
-    using open_contains_ball os p(2) by blast
-  then obtain ee where ee: "\<And>z. z \<in> path_image p \<Longrightarrow> 0 < ee z \<and> ball z (ee z) \<subseteq> s"
-    by metis
-  define cover where "cover = (\<lambda>z. ball z (ee z/3)) ` (path_image p)"
-  have "compact (path_image p)"
-    by (metis p(1) compact_path_image)
-  moreover have "path_image p \<subseteq> (\<Union>c\<in>path_image p. ball c (ee c / 3))"
-    using ee by auto
-  ultimately have "\<exists>D \<subseteq> cover. finite D \<and> path_image p \<subseteq> \<Union>D"
-    by (simp add: compact_eq_heine_borel cover_def)
-  then obtain D where D: "D \<subseteq> cover" "finite D" "path_image p \<subseteq> \<Union>D"
-    by blast
-  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"
-    apply (simp add: cover_def path_image_def image_comp)
-    apply (blast dest!: finite_subset_image [OF \<open>finite D\<close>])
-    done
-  then have kne: "k \<noteq> {}"
-    using D by auto
-  have pi: "\<And>i. i \<in> k \<Longrightarrow> p i \<in> path_image p"
-    using k  by (auto simp: path_image_def)
-  then have eepi: "\<And>i. i \<in> k \<Longrightarrow> 0 < ee((p i))"
-    by (metis ee)
-  define e where "e = Min((ee o p) ` k)"
-  have fin_eep: "finite ((ee o p) ` k)"
-    using k  by blast
-  have enz: "0 < e"
-    using ee k  by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi)
-  have "uniformly_continuous_on {0..1} p"
-    using p  by (simp add: path_def compact_uniformly_continuous)
-  then obtain d::real where d: "d>0"
-          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"
-    unfolding uniformly_continuous_on_def dist_norm real_norm_def
-    by (metis divide_pos_pos enz zero_less_numeral)
-  then obtain N::nat where N: "N>0" "inverse N < d"
-    using real_arch_inverse [of d]   by auto
-  { fix g h
-    assume g: "valid_path g" and gp: "\<forall>t\<in>{0..1}. cmod (g t - p t) < e / 3"
-       and h: "valid_path h" and hp: "\<forall>t\<in>{0..1}. cmod (h t - p t) < e / 3"
-       and joins: "linked_paths atends g h"
-    { fix t::real
-      assume t: "0 \<le> t" "t \<le> 1"
-      then obtain u where u: "u \<in> k" and ptu: "p t \<in> ball(p u) (ee(p u) / 3)"
-        using \<open>path_image p \<subseteq> \<Union>D\<close> D_eq by (force simp: path_image_def)
-      then have ele: "e \<le> ee (p u)" using fin_eep
-        by (simp add: e_def)
-      have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3"
-        using gp hp t by auto
-      with ele have "cmod (g t - p t) < ee (p u) / 3"
-                    "cmod (h t - p t) < ee (p u) / 3"
-        by linarith+
-      then have "g t \<in> ball(p u) (ee(p u))"  "h t \<in> ball(p u) (ee(p u))"
-        using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"]
-              norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u
-        by (force simp: dist_norm ball_def norm_minus_commute)+
-      then have "g t \<in> s" "h t \<in> s" using ee u k
-        by (auto simp: path_image_def ball_def)
-    }
-    then have ghs: "path_image g \<subseteq> s" "path_image h \<subseteq> s"
-      by (auto simp: path_image_def)
-    moreover
-    { fix f
-      assume fhols: "f holomorphic_on s"
-      then have fpa: "f contour_integrable_on g"  "f contour_integrable_on h"
-        using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple
-        by blast+
-      have contf: "continuous_on s f"
-        by (simp add: fhols holomorphic_on_imp_continuous_on)
-      { fix z
-        assume z: "z \<in> path_image p"
-        have "f holomorphic_on ball z (ee z)"
-          using fhols ee z holomorphic_on_subset by blast
-        then have "\<exists>ff. (\<forall>w \<in> ball z (ee z). (ff has_field_derivative f w) (at w))"
-          using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified]
-          by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball)
-      }
-      then obtain ff where ff:
-            "\<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)"
-        by metis
-      { fix n
-        assume n: "n \<le> N"
-        then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f =
-                   contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f"
-        proof (induct n)
-          case 0 show ?case by simp
-        next
-          case (Suc n)
-          obtain t where t: "t \<in> k" and "p (n/N) \<in> ball(p t) (ee(p t) / 3)"
-            using \<open>path_image p \<subseteq> \<Union>D\<close> [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems
-            by (force simp: path_image_def)
-          then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3"
-            by (simp add: dist_norm)
-          have e3le: "e/3 \<le> ee (p t) / 3"  using fin_eep t
-            by (simp add: e_def)
-          { fix x
-            assume x: "n/N \<le> x" "x \<le> (1 + n)/N"
-            then have nN01: "0 \<le> n/N" "(1 + n)/N \<le> 1"
-              using Suc.prems by auto
-            then have x01: "0 \<le> x" "x \<le> 1"
-              using x by linarith+
-            have "cmod (p t - p x)  < ee (p t) / 3 + e/3"
-              apply (rule norm_diff_triangle_less [OF ptu de])
-              using x N x01 Suc.prems
-              apply (auto simp: field_simps)
-              done
-            then have ptx: "cmod (p t - p x) < 2*ee (p t)/3"
-              using e3le eepi [OF t] by simp
-            have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 "
-              apply (rule norm_diff_triangle_less [OF ptx])
-              using gp x01 by (simp add: norm_minus_commute)
-            also have "... \<le> ee (p t)"
-              using e3le eepi [OF t] by simp
-            finally have gg: "cmod (p t - g x) < ee (p t)" .
-            have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 "
-              apply (rule norm_diff_triangle_less [OF ptx])
-              using hp x01 by (simp add: norm_minus_commute)
-            also have "... \<le> ee (p t)"
-              using e3le eepi [OF t] by simp
-            finally have "cmod (p t - g x) < ee (p t)"
-                         "cmod (p t - h x) < ee (p t)"
-              using gg by auto
-          } note ptgh_ee = this
-          have pi_hgn: "path_image (linepath (h (n/N)) (g (n/N))) \<subseteq> ball (p t) (ee (p t))"
-            using ptgh_ee [of "n/N"] Suc.prems
-            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
-          then have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \<subseteq> s"
-            using \<open>N>0\<close> Suc.prems
-            apply (simp add: path_image_join field_simps closed_segment_commute)
-            apply (erule order_trans)
-            apply (simp add: ee pi t)
-            done
-          have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N)))
-                  \<subseteq> ball (p t) (ee (p t))"
-            using ptgh_ee [of "(1+n)/N"] Suc.prems
-            by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"])
-          then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \<subseteq> s"
-            using \<open>N>0\<close> Suc.prems ee pi t
-            by (auto simp: Path_Connected.path_image_join field_simps)
-          have pi_subset_ball:
-                "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++
-                             subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N)))
-                 \<subseteq> ball (p t) (ee (p t))"
-            apply (intro subset_path_image_join pi_hgn pi_ghn')
-            using \<open>N>0\<close> Suc.prems
-            apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee)
-            done
-          have pi0: "(f has_contour_integral 0)
-                       (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++
-                        subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))"
-            apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"])
-            apply (metis ff open_ball at_within_open pi t)
-            apply (intro valid_path_join)
-            using Suc.prems pi_subset_ball apply (simp_all add: valid_path_subpath g h)
-            done
-          have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g"
-            using Suc.prems by (simp add: contour_integrable_subpath g fpa)
-          have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))"
-            using gh_n's
-            by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
-          have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))"
-            using gh_ns
-            by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf])
-          have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f +
-                     contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f +
-                     contour_integral (subpath ((Suc n) / N) (n/N) h) f +
-                     contour_integral (linepath (h (n/N)) (g (n/N))) f = 0"
-            using contour_integral_unique [OF pi0] Suc.prems
-            by (simp add: g h fpa valid_path_subpath contour_integrable_subpath
-                          fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc)
-          have *: "\<And>hn he hn' gn gd gn' hgn ghn gh0 ghn'.
-                    \<lbrakk>hn - gn = ghn - gh0;
-                     gd + ghn' + he + hgn = (0::complex);
-                     hn - he = hn'; gn + gd = gn'; hgn = -ghn\<rbrakk> \<Longrightarrow> hn' - gn' = ghn' - gh0"
-            by (auto simp: algebra_simps)
-          have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
-                contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f"
-            unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"]
-            using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath)
-          also have "... = contour_integral (subpath 0 ((Suc n) / N) h) f"
-            using Suc.prems by (simp add: contour_integral_subpath_combine h fpa)
-          finally have pi0_eq:
-               "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f =
-                contour_integral (subpath 0 ((Suc n) / N) h) f" .
-          show ?case
-            apply (rule * [OF Suc.hyps eq0 pi0_eq])
-            using Suc.prems
-            apply (simp_all add: g h fpa contour_integral_subpath_combine
-                     contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath
-                     continuous_on_subset [OF contf gh_ns])
-            done
-      qed
-      } note ind = this
-      have "contour_integral h f = contour_integral g f"
-        using ind [OF order_refl] N joins
-        by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm)
-    }
-    ultimately
-    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)"
-      by metis
-  } note * = this
-  show ?thesis
-    apply (rule_tac x="e/3" in exI)
-    apply (rule conjI)
-    using enz apply simp
-    apply (clarsimp simp only: ball_conj_distrib)
-    apply (rule *; assumption)
-    done
-qed
-
-
-lemma
-  assumes "open s" "path p" "path_image p \<subseteq> s"
-    shows contour_integral_nearby_ends:
-      "\<exists>d. 0 < d \<and>
-              (\<forall>g h. valid_path g \<and> valid_path h \<and>
-                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
-                    pathstart h = pathstart g \<and> pathfinish h = pathfinish g
-                    \<longrightarrow> 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))"
-    and contour_integral_nearby_loops:
-      "\<exists>d. 0 < d \<and>
-              (\<forall>g h. valid_path g \<and> valid_path h \<and>
-                    (\<forall>t \<in> {0..1}. norm(g t - p t) < d \<and> norm(h t - p t) < d) \<and>
-                    pathfinish g = pathstart g \<and> pathfinish h = pathstart h
-                    \<longrightarrow> 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))"
-  using contour_integral_nearby [OF assms, where atends=True]
-  using contour_integral_nearby [OF assms, where atends=False]
-  unfolding linked_paths_def by simp_all
-
-corollary differentiable_polynomial_function:
-  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "polynomial_function p \<Longrightarrow> p differentiable_on s"
-by (meson has_vector_derivative_polynomial_function differentiable_at_imp_differentiable_on differentiable_def has_vector_derivative_def)
-
-lemma C1_differentiable_polynomial_function:
-  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "polynomial_function p \<Longrightarrow> p C1_differentiable_on s"
-  by (metis continuous_on_polymonial_function C1_differentiable_on_def  has_vector_derivative_polynomial_function)
-
-lemma valid_path_polynomial_function:
-  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "polynomial_function p \<Longrightarrow> valid_path p"
-by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function)
-
-lemma valid_path_subpath_trivial [simp]:
-    fixes g :: "real \<Rightarrow> 'a::euclidean_space"
-    shows "z \<noteq> g x \<Longrightarrow> valid_path (subpath x x g)"
-  by (simp add: subpath_def valid_path_polynomial_function)
-
-lemma contour_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(contour_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> contour_integral g f = contour_integral p f"
-  using contour_integral_nearby_ends [OF s \<open>path g\<close> 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 contour_integrable_on p \<and> valid_path p"
-    using p s
-    by (blast intro: valid_path_polynomial_function contour_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 (contour_integral g f) \<le> L * B"
-    apply (simp add: pi [OF f])
-    apply (simp add: contour_integral_integral)
-    apply (rule order_trans [OF integral_norm_bound_integral])
-    apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult)
-    done
-} then
-show ?thesis
-  by (force simp: L contour_integral_integral)
-qed
-
-subsection\<open>Constancy of a function from a connected set into a finite, disconnected or discrete set\<close>
-
-text\<open>Still missing: versions for a set that is smaller than R, or countable.\<close>
-
-lemma continuous_disconnected_range_constant:
-  assumes s: "connected s"
-      and conf: "continuous_on s f"
-      and fim: "f ` s \<subseteq> t"
-      and cct: "\<And>y. y \<in> t \<Longrightarrow> connected_component_set t y = {y}"
-    shows "\<exists>a. \<forall>x \<in> s. f x = a"
-proof (cases "s = {}")
-  case True then show ?thesis by force
-next
-  case False
-  { fix x assume "x \<in> s"
-    then have "f ` s \<subseteq> {f x}"
-    by (metis connected_continuous_image conf connected_component_maximal fim image_subset_iff rev_image_eqI s cct)
-  }
-  with False show ?thesis
-    by blast
-qed
-
-lemma discrete_subset_disconnected:
-  fixes s :: "'a::topological_space set"
-  fixes t :: "'b::real_normed_vector set"
-  assumes conf: "continuous_on s f"
-      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)"
-   shows "f ` s \<subseteq> {y. connected_component_set (f ` s) y = {y}}"
-proof -
-  { fix x assume x: "x \<in> s"
-    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)"
-      using conf no [OF x] by auto
-    then have e2: "0 \<le> e / 2"
-      by simp
-    have "f y = f x" if "y \<in> s" and ccs: "f y \<in> connected_component_set (f ` s) (f x)" for y
-      apply (rule ccontr)
-      using connected_closed [of "connected_component_set (f ` s) (f x)"] \<open>e>0\<close>
-      apply (simp add: del: ex_simps)
-      apply (drule spec [where x="cball (f x) (e / 2)"])
-      apply (drule spec [where x="- ball(f x) e"])
-      apply (auto simp: dist_norm open_closed [symmetric] simp del: le_divide_eq_numeral1 dest!: connected_component_in)
-        apply (metis diff_self e2 ele norm_minus_commute norm_zero not_less)
-       using centre_in_cball connected_component_refl_eq e2 x apply blast
-      using ccs
-      apply (force simp: cball_def dist_norm norm_minus_commute dest: ele [OF \<open>y \<in> s\<close>])
-      done
-    moreover have "connected_component_set (f ` s) (f x) \<subseteq> f ` s"
-      by (auto simp: connected_component_in)
-    ultimately have "connected_component_set (f ` s) (f x) = {f x}"
-      by (auto simp: x)
-  }
-  with assms show ?thesis
-    by blast
-qed
-
-lemma finite_implies_discrete:
-  fixes s :: "'a::topological_space set"
-  assumes "finite (f ` s)"
-  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))"
-proof -
-  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
-  proof (cases "f ` s - {f x} = {}")
-    case True
-    with zero_less_numeral show ?thesis
-      by (fastforce simp add: Set.image_subset_iff cong: conj_cong)
-  next
-    case False
-    then obtain z where z: "z \<in> s" "f z \<noteq> f x"
-      by blast
-    have finn: "finite {norm (z - f x) |z. z \<in> f ` s - {f x}}"
-      using assms by simp
-    then have *: "0 < Inf{norm(z - f x) | z. z \<in> f ` s - {f x}}"
-      apply (rule finite_imp_less_Inf)
-      using z apply force+
-      done
-    show ?thesis
-      by (force intro!: * cInf_le_finite [OF finn])
-  qed
-  with assms show ?thesis
-    by blast
-qed
-
-text\<open>This proof requires the existence of two separate values of the range type.\<close>
-lemma finite_range_constant_imp_connected:
-  assumes "\<And>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
-              \<lbrakk>continuous_on s f; finite(f ` s)\<rbrakk> \<Longrightarrow> \<exists>a. \<forall>x \<in> s. f x = a"
-    shows "connected s"
-proof -
-  { fix t u
-    assume clt: "closedin (subtopology euclidean s) t"
-       and clu: "closedin (subtopology euclidean s) u"
-       and tue: "t \<inter> u = {}" and tus: "t \<union> u = s"
-    have conif: "continuous_on s (\<lambda>x. if x \<in> t then 0 else 1)"
-      apply (subst tus [symmetric])
-      apply (rule continuous_on_cases_local)
-      using clt clu tue
-      apply (auto simp: tus continuous_on_const)
-      done
-    have fi: "finite ((\<lambda>x. if x \<in> t then 0 else 1) ` s)"
-      by (rule finite_subset [of _ "{0,1}"]) auto
-    have "t = {} \<or> u = {}"
-      using assms [OF conif fi] tus [symmetric]
-      by (auto simp: Ball_def) (metis IntI empty_iff one_neq_zero tue)
-  }
-  then show ?thesis
-    by (simp add: connected_closedin_eq)
-qed
-
-lemma continuous_disconnected_range_constant_eq:
-      "(connected s \<longleftrightarrow>
-           (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
-            \<forall>t. continuous_on s f \<and> f ` s \<subseteq> t \<and> (\<forall>y \<in> t. connected_component_set t y = {y})
-            \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis1)
-  and continuous_discrete_range_constant_eq:
-      "(connected s \<longleftrightarrow>
-         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
-          continuous_on s f \<and>
-          (\<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)))
-          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis2)
-  and continuous_finite_range_constant_eq:
-      "(connected s \<longleftrightarrow>
-         (\<forall>f::'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1.
-          continuous_on s f \<and> finite (f ` s)
-          \<longrightarrow> (\<exists>a::'b. \<forall>x \<in> s. f x = a)))" (is ?thesis3)
-proof -
-  have *: "\<And>s t u v. \<lbrakk>s \<Longrightarrow> t; t \<Longrightarrow> u; u \<Longrightarrow> v; v \<Longrightarrow> s\<rbrakk>
-    \<Longrightarrow> (s \<longleftrightarrow> t) \<and> (s \<longleftrightarrow> u) \<and> (s \<longleftrightarrow> v)"
-    by blast
-  have "?thesis1 \<and> ?thesis2 \<and> ?thesis3"
-    apply (rule *)
-    using continuous_disconnected_range_constant apply metis
-    apply clarify
-    apply (frule discrete_subset_disconnected; blast)
-    apply (blast dest: finite_implies_discrete)
-    apply (blast intro!: finite_range_constant_imp_connected)
-    done
-  then show ?thesis1 ?thesis2 ?thesis3
-    by blast+
-qed
-
-lemma continuous_discrete_range_constant:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
-  assumes s: "connected s"
-      and "continuous_on s f"
-      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)"
-    shows "\<exists>a. \<forall>x \<in> s. f x = a"
-  using continuous_discrete_range_constant_eq [THEN iffD1, OF s] assms
-  by blast
-
-lemma continuous_finite_range_constant:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::real_normed_algebra_1"
-  assumes "connected s"
-      and "continuous_on s f"
-      and "finite (f ` s)"
-    shows "\<exists>a. \<forall>x \<in> s. f x = a"
-  using assms continuous_finite_range_constant_eq
-  by blast
-
-
-text\<open>We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\<close>
-
-subsection\<open>Winding Numbers\<close>
-
-definition winding_number:: "[real \<Rightarrow> complex, complex] \<Rightarrow> complex" where
-  "winding_number \<gamma> z \<equiv>
-    @n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
-                    pathstart p = pathstart \<gamma> \<and>
-                    pathfinish p = pathfinish \<gamma> \<and>
-                    (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
-                    contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-
-lemma winding_number:
-  assumes "path \<gamma>" "z \<notin> path_image \<gamma>" "0 < e"
-    shows "\<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
-               pathstart p = pathstart \<gamma> \<and>
-               pathfinish p = pathfinish \<gamma> \<and>
-               (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
-               contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * winding_number \<gamma> z"
-proof -
-  have "path_image \<gamma> \<subseteq> UNIV - {z}"
-    using assms by blast
-  then obtain d
-    where d: "d>0"
-      and pi_eq: "\<And>h1 h2. valid_path h1 \<and> valid_path h2 \<and>
-                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d) \<and>
-                    pathstart h2 = pathstart h1 \<and> pathfinish h2 = pathfinish h1 \<longrightarrow>
-                      path_image h1 \<subseteq> UNIV - {z} \<and> path_image h2 \<subseteq> UNIV - {z} \<and>
-                      (\<forall>f. f holomorphic_on UNIV - {z} \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
-    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms by (auto simp: open_delete)
-  then obtain h where h: "polynomial_function h \<and> pathstart h = pathstart \<gamma> \<and> pathfinish h = pathfinish \<gamma> \<and>
-                          (\<forall>t \<in> {0..1}. norm(h t - \<gamma> t) < d/2)"
-    using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "d/2"] d by auto
-  define nn where "nn = 1/(2* pi*\<i>) * contour_integral h (\<lambda>w. 1/(w - z))"
-  have "\<exists>n. \<forall>e > 0. \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
-                        pathstart p = pathstart \<gamma> \<and>  pathfinish p = pathfinish \<gamma> \<and>
-                        (\<forall>t \<in> {0..1}. norm(\<gamma> t - p t) < e) \<and>
-                        contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-                    (is "\<exists>n. \<forall>e > 0. ?PP e n")
-    proof (rule_tac x=nn in exI, clarify)
-      fix e::real
-      assume e: "e>0"
-      obtain p where p: "polynomial_function p \<and>
-            pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and> (\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < min e (d / 2))"
-        using path_approx_polynomial_function [OF \<open>path \<gamma>\<close>, of "min e (d/2)"] d \<open>0<e\<close> by auto
-      have "(\<lambda>w. 1 / (w - z)) holomorphic_on UNIV - {z}"
-        by (auto simp: intro!: holomorphic_intros)
-      then show "?PP e nn"
-        apply (rule_tac x=p in exI)
-        using pi_eq [of h p] h p d
-        apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def)
-        done
-    qed
-  then show ?thesis
-    unfolding winding_number_def
-    apply (rule someI2_ex)
-    apply (blast intro: \<open>0<e\<close>)
-    done
-qed
-
-lemma winding_number_unique:
-  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
-      and pi:
-        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
-                          pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
-                          (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
-                          contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-   shows "winding_number \<gamma> z = n"
-proof -
-  have "path_image \<gamma> \<subseteq> UNIV - {z}"
-    using assms by blast
-  then obtain e
-    where e: "e>0"
-      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
-                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
-                    pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
-                    contour_integral h2 f = contour_integral h1 f"
-    using contour_integral_nearby_ends [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
-  obtain p where p:
-     "valid_path p \<and> z \<notin> path_image p \<and>
-      pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
-      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
-      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-    using pi [OF e] by blast
-  obtain q where q:
-     "valid_path q \<and> z \<notin> path_image q \<and>
-      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
-      (\<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"
-    using winding_number [OF \<gamma> e] by blast
-  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
-    using p by auto
-  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
-    apply (rule pi_eq)
-    using p q
-    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
-  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
-    using q by auto
-  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
-  then show ?thesis
-    by simp
-qed
-
-lemma winding_number_unique_loop:
-  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
-      and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-      and pi:
-        "\<And>e. e>0 \<Longrightarrow> \<exists>p. valid_path p \<and> z \<notin> path_image p \<and>
-                           pathfinish p = pathstart p \<and>
-                           (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
-                           contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-   shows "winding_number \<gamma> z = n"
-proof -
-  have "path_image \<gamma> \<subseteq> UNIV - {z}"
-    using assms by blast
-  then obtain e
-    where e: "e>0"
-      and pi_eq: "\<And>h1 h2 f. \<lbrakk>valid_path h1; valid_path h2;
-                    (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < e \<and> cmod (h2 t - \<gamma> t) < e);
-                    pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\<rbrakk> \<Longrightarrow>
-                    contour_integral h2 f = contour_integral h1 f"
-    using contour_integral_nearby_loops [of "UNIV - {z}" \<gamma>] assms  by (auto simp: open_delete)
-  obtain p where p:
-     "valid_path p \<and> z \<notin> path_image p \<and>
-      pathfinish p = pathstart p \<and>
-      (\<forall>t \<in> {0..1}. norm (\<gamma> t - p t) < e) \<and>
-      contour_integral p (\<lambda>w. 1/(w - z)) = 2 * pi * \<i> * n"
-    using pi [OF e] by blast
-  obtain q where q:
-     "valid_path q \<and> z \<notin> path_image q \<and>
-      pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma> \<and>
-      (\<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"
-    using winding_number [OF \<gamma> e] by blast
-  have "2 * complex_of_real pi * \<i> * n = contour_integral p (\<lambda>w. 1 / (w - z))"
-    using p by auto
-  also have "... = contour_integral q (\<lambda>w. 1 / (w - z))"
-    apply (rule pi_eq)
-    using p q loop
-    by (auto simp: valid_path_polynomial_function norm_minus_commute intro!: holomorphic_intros)
-  also have "... = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z"
-    using q by auto
-  finally have "2 * complex_of_real pi * \<i> * n = 2 * complex_of_real pi * \<i> * winding_number \<gamma> z" .
-  then show ?thesis
-    by simp
-qed
-
-lemma winding_number_valid_path:
-  assumes "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
-    shows "winding_number \<gamma> z = 1/(2*pi*\<i>) * contour_integral \<gamma> (\<lambda>w. 1/(w - z))"
-  using assms by (auto simp: valid_path_imp_path intro!: winding_number_unique)
-
-lemma has_contour_integral_winding_number:
-  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
-    shows "((\<lambda>w. 1/(w - z)) has_contour_integral (2*pi*\<i>*winding_number \<gamma> z)) \<gamma>"
-by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms)
-
-lemma winding_number_trivial [simp]: "z \<noteq> a \<Longrightarrow> winding_number(linepath a a) z = 0"
-  by (simp add: winding_number_valid_path)
-
-lemma winding_number_subpath_trivial [simp]: "z \<noteq> g x \<Longrightarrow> winding_number (subpath x x g) z = 0"
-  by (simp add: path_image_subpath winding_number_valid_path)
-
-lemma winding_number_join:
-  assumes g1: "path g1" "z \<notin> path_image g1"
-      and g2: "path g2" "z \<notin> path_image g2"
-      and "pathfinish g1 = pathstart g2"
-    shows "winding_number(g1 +++ g2) z = winding_number g1 z + winding_number g2 z"
-  apply (rule winding_number_unique)
-  using assms apply (simp_all add: not_in_path_image_join)
-  apply (frule winding_number [OF g2])
-  apply (frule winding_number [OF g1], clarify)
-  apply (rename_tac p2 p1)
-  apply (rule_tac x="p1+++p2" in exI)
-  apply (simp add: not_in_path_image_join contour_integrable_inversediff algebra_simps)
-  apply (auto simp: joinpaths_def)
-  done
-
-lemma winding_number_reversepath:
-  assumes "path \<gamma>" "z \<notin> path_image \<gamma>"
-    shows "winding_number(reversepath \<gamma>) z = - (winding_number \<gamma> z)"
-  apply (rule winding_number_unique)
-  using assms
-  apply simp_all
-  apply (frule winding_number [OF assms], clarify)
-  apply (rule_tac x="reversepath p" in exI)
-  apply (simp add: contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse)
-  apply (auto simp: reversepath_def)
-  done
-
-lemma winding_number_shiftpath:
-  assumes \<gamma>: "path \<gamma>" "z \<notin> path_image \<gamma>"
-      and "pathfinish \<gamma> = pathstart \<gamma>" "a \<in> {0..1}"
-    shows "winding_number(shiftpath a \<gamma>) z = winding_number \<gamma> z"
-  apply (rule winding_number_unique_loop)
-  using assms
-  apply (simp_all add: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath)
-  apply (frule winding_number [OF \<gamma>], clarify)
-  apply (rule_tac x="shiftpath a p" in exI)
-  apply (simp add: contour_integral_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath valid_path_shiftpath)
-  apply (auto simp: shiftpath_def)
-  done
-
-lemma winding_number_split_linepath:
-  assumes "c \<in> closed_segment a b" "z \<notin> closed_segment a b"
-    shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z"
-proof -
-  have "z \<notin> closed_segment a c" "z \<notin> closed_segment c b"
-    using assms  apply (meson convex_contains_segment convex_segment ends_in_segment(1) subsetCE)
-    using assms  by (meson convex_contains_segment convex_segment ends_in_segment(2) subsetCE)
-  then show ?thesis
-    using assms
-    by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps)
-qed
-
-lemma winding_number_cong:
-   "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> p t = q t) \<Longrightarrow> winding_number p z = winding_number q z"
-  by (simp add: winding_number_def pathstart_def pathfinish_def)
-
-lemma winding_number_offset: "winding_number p z = winding_number (\<lambda>w. p w - z) 0"
-  apply (simp add: winding_number_def contour_integral_integral path_image_def valid_path_def pathstart_def pathfinish_def)
-  apply (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe)
-  apply (rename_tac g)
-  apply (rule_tac x="\<lambda>t. g t - z" in exI)
-  apply (force simp: vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise)
-  apply (rename_tac g)
-  apply (rule_tac x="\<lambda>t. g t + z" in exI)
-  apply (simp add: piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise)
-  apply (force simp: algebra_simps)
-  done
-
-(* A combined theorem deducing several things piecewise.*)
-lemma winding_number_join_pos_combined:
-     "\<lbrakk>valid_path \<gamma>1; z \<notin> path_image \<gamma>1; 0 < Re(winding_number \<gamma>1 z);
-       valid_path \<gamma>2; z \<notin> path_image \<gamma>2; 0 < Re(winding_number \<gamma>2 z); pathfinish \<gamma>1 = pathstart \<gamma>2\<rbrakk>
-      \<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)"
-  by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path)
-
-
-(* Useful sufficient conditions for the winding number to be positive etc.*)
-
-lemma Re_winding_number:
-    "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>\<rbrakk>
-     \<Longrightarrow> Re(winding_number \<gamma> z) = Im(contour_integral \<gamma> (\<lambda>w. 1/(w - z))) / (2*pi)"
-by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square)
-
-lemma winding_number_pos_le:
-  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
-      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> 0 \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
-    shows "0 \<le> Re(winding_number \<gamma> z)"
-proof -
-  have *: "0 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))" if x: "0 < x" "x < 1" for x
-    using ge by (simp add: Complex.Im_divide algebra_simps x)
-  show ?thesis
-    apply (simp add: Re_winding_number [OF \<gamma>] field_simps)
-    apply (rule has_integral_component_nonneg
-             [of \<i> "\<lambda>x. if x \<in> {0<..<1}
-                         then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else 0", simplified])
-      prefer 3 apply (force simp: *)
-     apply (simp add: Basis_complex_def)
-    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)"])
-    apply simp
-    apply (simp only: box_real)
-    apply (subst has_contour_integral [symmetric])
-    using \<gamma>
-    apply (simp add: contour_integrable_inversediff has_contour_integral_integral)
-    done
-qed
-
-lemma winding_number_pos_lt_lemma:
-  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
-      and e: "0 < e"
-      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
-    shows "0 < Re(winding_number \<gamma> z)"
-proof -
-  have "e \<le> Im (contour_integral \<gamma> (\<lambda>w. 1 / (w - z)))"
-    apply (rule has_integral_component_le
-             [of \<i> "\<lambda>x. \<i>*e" "\<i>*e" "{0..1}"
-                    "\<lambda>x. if x \<in> {0<..<1} then 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x) else \<i>*e"
-                    "contour_integral \<gamma> (\<lambda>w. 1/(w - z))", simplified])
-    using e
-    apply (simp_all add: Basis_complex_def)
-    using has_integral_const_real [of _ 0 1] apply force
-    apply (rule has_integral_spike_interior [of 0 1 _ "\<lambda>x. 1/(\<gamma> x - z) * vector_derivative \<gamma> (at x)", simplified box_real])
-    apply simp
-    apply (subst has_contour_integral [symmetric])
-    using \<gamma>
-    apply (simp_all add: contour_integrable_inversediff has_contour_integral_integral ge)
-    done
-  with e show ?thesis
-    by (simp add: Re_winding_number [OF \<gamma>] field_simps)
-qed
-
-lemma winding_number_pos_lt:
-  assumes \<gamma>: "valid_path \<gamma>" "z \<notin> path_image \<gamma>"
-      and e: "0 < e"
-      and ge: "\<And>x. \<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> e \<le> Im (vector_derivative \<gamma> (at x) * cnj(\<gamma> x - z))"
-    shows "0 < Re (winding_number \<gamma> z)"
-proof -
-  have bm: "bounded ((\<lambda>w. w - z) ` (path_image \<gamma>))"
-    using bounded_translation [of _ "-z"] \<gamma> by (simp add: bounded_valid_path_image)
-  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"
-    using bounded_pos [THEN iffD1, OF bm] by blast
-  { fix x::real  assume x: "0 < x" "x < 1"
-    then have B2: "cmod (\<gamma> x - z)^2 \<le> B^2" using Bno [of "\<gamma> x - z"]
-      by (simp add: path_image_def power2_eq_square mult_mono')
-    with x have "\<gamma> x \<noteq> z" using \<gamma>
-      using path_image_def by fastforce
-    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) * cnj (\<gamma> x - z)) / (cmod (\<gamma> x - z))\<^sup>2"
-      using B ge [OF x] B2 e
-      apply (rule_tac y="e / (cmod (\<gamma> x - z))\<^sup>2" in order_trans)
-      apply (auto simp: divide_left_mono divide_right_mono)
-      done
-    then have "e / B\<^sup>2 \<le> Im (vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
-      by (simp add: Im_divide_Reals complex_div_cnj [of _ "\<gamma> x - z" for x] del: complex_cnj_diff times_complex.sel)
-  } note * = this
-  show ?thesis
-    using e B by (simp add: * winding_number_pos_lt_lemma [OF \<gamma>, of "e/B^2"])
-qed
-
-subsection\<open>The winding number is an integer\<close>
-
-text\<open>Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1,
-     Also on page 134 of Serge Lang's book with the name title, etc.\<close>
-
-lemma exp_fg:
-  fixes z::complex
-  assumes g: "(g has_vector_derivative g') (at x within s)"
-      and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)"
-      and z: "g x \<noteq> z"
-    shows "((\<lambda>x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)"
-proof -
-  have *: "(exp o (\<lambda>x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)"
-    using assms unfolding has_vector_derivative_def scaleR_conv_of_real
-    by (auto intro!: derivative_eq_intros)
-  show ?thesis
-    apply (rule has_vector_derivative_eq_rhs)
-    apply (rule bounded_bilinear.has_vector_derivative [OF bounded_bilinear_mult])
-    using z
-    apply (auto simp: intro!: derivative_eq_intros * [unfolded o_def] g)
-    done
-qed
-
-lemma winding_number_exp_integral:
-  fixes z::complex
-  assumes \<gamma>: "\<gamma> piecewise_C1_differentiable_on {a..b}"
-      and ab: "a \<le> b"
-      and z: "z \<notin> \<gamma> ` {a..b}"
-    shows "(\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)) integrable_on {a..b}"
-          (is "?thesis1")
-          "exp (- (integral {a..b} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))) * (\<gamma> b - z) = \<gamma> a - z"
-          (is "?thesis2")
-proof -
-  let ?D\<gamma> = "\<lambda>x. vector_derivative \<gamma> (at x)"
-  have [simp]: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<gamma> x \<noteq> z"
-    using z by force
-  have cong: "continuous_on {a..b} \<gamma>"
-    using \<gamma> by (simp add: piecewise_C1_differentiable_on_def)
-  obtain k where fink: "finite k" and g_C1_diff: "\<gamma> C1_differentiable_on ({a..b} - k)"
-    using \<gamma> by (force simp: piecewise_C1_differentiable_on_def)
-  have o: "open ({a<..<b} - k)"
-    using \<open>finite k\<close> by (simp add: finite_imp_closed open_Diff)
-  moreover have "{a<..<b} - k \<subseteq> {a..b} - k"
-    by force
-  ultimately have g_diff_at: "\<And>x. \<lbrakk>x \<notin> k; x \<in> {a<..<b}\<rbrakk> \<Longrightarrow> \<gamma> differentiable at x"
-    by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def differentiable_within_open)
-  { fix w
-    assume "w \<noteq> z"
-    have "continuous_on (ball w (cmod (w - z))) (\<lambda>w. 1 / (w - z))"
-      by (auto simp: dist_norm intro!: continuous_intros)
-    moreover have "\<And>x. cmod (w - x) < cmod (w - z) \<Longrightarrow> \<exists>f'. ((\<lambda>w. 1 / (w - z)) has_field_derivative f') (at x)"
-      by (auto simp: intro!: derivative_eq_intros)
-    ultimately have "\<exists>h. \<forall>y. norm(y - w) < norm(w - z) \<longrightarrow> (h has_field_derivative 1/(y - z)) (at y)"
-      using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\<lambda>w. 1/(w - z)"]
-      by (simp add: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute)
-  }
-  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)"
-    by meson
-  have exy: "\<exists>y. ((\<lambda>x. inverse (\<gamma> x - z) * ?D\<gamma> x) has_integral y) {a..b}"
-    unfolding integrable_on_def [symmetric]
-    apply (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \<gamma>], of "-{z}"])
-    apply (rename_tac w)
-    apply (rule_tac x="norm(w - z)" in exI)
-    apply (simp_all add: inverse_eq_divide)
-    apply (metis has_field_derivative_at_within h)
-    done
-  have vg_int: "(\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)) integrable_on {a..b}"
-    unfolding box_real [symmetric] divide_inverse_commute
-    by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab)
-  with ab show ?thesis1
-    by (simp add: divide_inverse_commute integral_def integrable_on_def)
-  { fix t
-    assume t: "t \<in> {a..b}"
-    have cball: "continuous_on (ball (\<gamma> t) (dist (\<gamma> t) z)) (\<lambda>x. inverse (x - z))"
-        using z by (auto intro!: continuous_intros simp: dist_norm)
-    have icd: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow> (\<lambda>w. inverse (w - z)) field_differentiable at x"
-      unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros)
-    obtain h where h: "\<And>x. cmod (\<gamma> t - x) < cmod (\<gamma> t - z) \<Longrightarrow>
-                       (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\<gamma> t - y) < cmod (\<gamma> t - z)})"
-      using holomorphic_convex_primitive [where f = "\<lambda>w. inverse(w - z)", OF convex_ball finite.emptyI cball icd]
-      by simp (auto simp: ball_def dist_norm that)
-    { fix x D
-      assume x: "x \<notin> k" "a < x" "x < b"
-      then have "x \<in> interior ({a..b} - k)"
-        using open_subset_interior [OF o] by fastforce
-      then have con: "isCont (\<lambda>x. ?D\<gamma> x) x"
-        using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior)
-      then have con_vd: "continuous (at x within {a..b}) (\<lambda>x. ?D\<gamma> x)"
-        by (rule continuous_at_imp_continuous_within)
-      have gdx: "\<gamma> differentiable at x"
-        using x by (simp add: g_diff_at)
-      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))
-          (at x within {a..b})"
-        using x gdx t
-        apply (clarsimp simp add: differentiable_iff_scaleR)
-        apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_within)
-        apply (simp_all add: has_vector_derivative_def [symmetric])
-        apply (rule has_vector_derivative_eq_rhs [OF integral_has_vector_derivative_continuous_at])
-        apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+
-        done
-      } note * = this
-    have "exp (- (integral {a..t} (\<lambda>x. ?D\<gamma> x / (\<gamma> x - z)))) * (\<gamma> t - z) =\<gamma> a - z"
-      apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \<union> k" a b])
-      using t
-      apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous [OF vg_int]  simp add: ab)+
-      done
-   }
-  with ab show ?thesis2
-    by (simp add: divide_inverse_commute integral_def)
-qed
-
-corollary winding_number_exp_2pi:
-    "\<lbrakk>path p; z \<notin> path_image p\<rbrakk>
-     \<Longrightarrow> pathfinish p - z = exp (2 * pi * \<i> * winding_number p z) * (pathstart p - z)"
-using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def
-  by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus)
-
-
-subsection\<open>The version with complex integers and equality\<close>
-
-lemma integer_winding_number_eq:
-  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
-  shows "winding_number \<gamma> z \<in> \<int> \<longleftrightarrow> pathfinish \<gamma> = pathstart \<gamma>"
-proof -
-  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)"
-      by (simp add: field_simps) algebra
-  obtain p where p: "valid_path p" "z \<notin> path_image p"
-                    "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
-                    "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
-    using winding_number [OF assms, of 1] by auto
-  have [simp]: "(winding_number \<gamma> z \<in> \<int>) = (exp (contour_integral p (\<lambda>w. 1 / (w - z))) = 1)"
-      using p by (simp add: exp_eq_1 complex_is_Int_iff)
-  have "winding_number p z \<in> \<int> \<longleftrightarrow> pathfinish p = pathstart p"
-    using p z
-    apply (simp add: winding_number_valid_path valid_path_def path_image_def pathstart_def pathfinish_def)
-    using winding_number_exp_integral(2) [of p 0 1 z]
-    apply (simp add: field_simps contour_integral_integral exp_minus)
-    apply (rule *)
-    apply (auto simp: path_image_def field_simps)
-    done
-  then show ?thesis using p
-    by (auto simp: winding_number_valid_path)
-qed
-
-theorem integer_winding_number:
-  "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> path_image \<gamma>\<rbrakk> \<Longrightarrow> winding_number \<gamma> z \<in> \<int>"
-by (metis integer_winding_number_eq)
-
-
-text\<open>If the winding number's magnitude is at least one, then the path must contain points in every direction.*)
-   We can thus bound the winding number of a path that doesn't intersect a given ray. \<close>
-
-lemma winding_number_pos_meets:
-  fixes z::complex
-  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and 1: "Re (winding_number \<gamma> z) \<ge> 1"
-      and w: "w \<noteq> z"
-  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
-proof -
-  have [simp]: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<gamma> x \<noteq> z"
-    using z by (auto simp: path_image_def)
-  have [simp]: "z \<notin> \<gamma> ` {0..1}"
-    using path_image_def z by auto
-  have gpd: "\<gamma> piecewise_C1_differentiable_on {0..1}"
-    using \<gamma> valid_path_def by blast
-  define r where "r = (w - z) / (\<gamma> 0 - z)"
-  have [simp]: "r \<noteq> 0"
-    using w z by (auto simp: r_def)
-  have "Arg r \<le> 2*pi"
-    by (simp add: Arg less_eq_real_def)
-  also have "... \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))"
-    using 1
-    apply (simp add: winding_number_valid_path [OF \<gamma> z] Cauchy_Integral_Thm.contour_integral_integral)
-    apply (simp add: Complex.Re_divide field_simps power2_eq_square)
-    done
-  finally have "Arg r \<le> Im (integral {0..1} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z)))" .
-  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"
-    apply (simp add:)
-    apply (rule Topological_Spaces.IVT')
-    apply (simp_all add: Complex_Transcendental.Arg_ge_0)
-    apply (intro continuous_intros indefinite_integral_continuous winding_number_exp_integral [OF gpd]; simp)
-    done
-  then obtain t where t:     "t \<in> {0..1}"
-                  and eqArg: "Im (integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x)/(\<gamma> x - z))) = Arg r"
-    by blast
-  define i where "i = integral {0..t} (\<lambda>x. vector_derivative \<gamma> (at x) / (\<gamma> x - z))"
-  have iArg: "Arg r = Im i"
-    using eqArg by (simp add: i_def)
-  have gpdt: "\<gamma> piecewise_C1_differentiable_on {0..t}"
-    by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t)
-  have "exp (- i) * (\<gamma> t - z) = \<gamma> 0 - z"
-    unfolding i_def
-    apply (rule winding_number_exp_integral [OF gpdt])
-    using t z unfolding path_image_def
-    apply force+
-    done
-  then have *: "\<gamma> t - z = exp i * (\<gamma> 0 - z)"
-    by (simp add: exp_minus field_simps)
-  then have "(w - z) = r * (\<gamma> 0 - z)"
-    by (simp add: r_def)
-  then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \<gamma> t"
-    apply (simp add:)
-    apply (subst Complex_Transcendental.Arg_eq [of r])
-    apply (simp add: iArg)
-    using *
-    apply (simp add: exp_eq_polar field_simps)
-    done
-  with t show ?thesis
-    by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def)
-qed
-
-lemma winding_number_big_meets:
-  fixes z::complex
-  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "\<bar>Re (winding_number \<gamma> z)\<bar> \<ge> 1"
-      and w: "w \<noteq> z"
-  shows "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image \<gamma>"
-proof -
-  { assume "Re (winding_number \<gamma> z) \<le> - 1"
-    then have "Re (winding_number (reversepath \<gamma>) z) \<ge> 1"
-      by (simp add: \<gamma> valid_path_imp_path winding_number_reversepath z)
-    moreover have "valid_path (reversepath \<gamma>)"
-      using \<gamma> valid_path_imp_reverse by auto
-    moreover have "z \<notin> path_image (reversepath \<gamma>)"
-      by (simp add: z)
-    ultimately have "\<exists>a::real. 0 < a \<and> z + a*(w - z) \<in> path_image (reversepath \<gamma>)"
-      using winding_number_pos_meets w by blast
-    then have ?thesis
-      by simp
-  }
-  then show ?thesis
-    using assms
-    by (simp add: Groups.abs_if_class.abs_if winding_number_pos_meets split: if_split_asm)
-qed
-
-lemma winding_number_less_1:
-  fixes z::complex
-  shows
-  "\<lbrakk>valid_path \<gamma>; z \<notin> path_image \<gamma>; w \<noteq> z;
-    \<And>a::real. 0 < a \<Longrightarrow> z + a*(w - z) \<notin> path_image \<gamma>\<rbrakk>
-   \<Longrightarrow> \<bar>Re(winding_number \<gamma> z)\<bar> < 1"
-   by (auto simp: not_less dest: winding_number_big_meets)
-
-text\<open>One way of proving that WN=1 for a loop.\<close>
-lemma winding_number_eq_1:
-  fixes z::complex
-  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-      and 0: "0 < Re(winding_number \<gamma> z)" and 2: "Re(winding_number \<gamma> z) < 2"
-  shows "winding_number \<gamma> z = 1"
-proof -
-  have "winding_number \<gamma> z \<in> Ints"
-    by (simp add: \<gamma> integer_winding_number loop valid_path_imp_path z)
-  then show ?thesis
-    using 0 2 by (auto simp: Ints_def)
-qed
-
-
-subsection\<open>Continuity of winding number and invariance on connected sets.\<close>
-
-lemma continuous_at_winding_number:
-  fixes z::complex
-  assumes \<gamma>: "path \<gamma>" and z: "z \<notin> path_image \<gamma>"
-  shows "continuous (at z) (winding_number \<gamma>)"
-proof -
-  obtain e where "e>0" and cbg: "cball z e \<subseteq> - path_image \<gamma>"
-    using open_contains_cball [of "- path_image \<gamma>"]  z
-    by (force simp: closed_def [symmetric] closed_path_image [OF \<gamma>])
-  then have ppag: "path_image \<gamma> \<subseteq> - cball z (e/2)"
-    by (force simp: cball_def dist_norm)
-  have oc: "open (- cball z (e / 2))"
-    by (simp add: closed_def [symmetric])
-  obtain d where "d>0" and pi_eq:
-    "\<And>h1 h2. \<lbrakk>valid_path h1; valid_path h2;
-              (\<forall>t\<in>{0..1}. cmod (h1 t - \<gamma> t) < d \<and> cmod (h2 t - \<gamma> t) < d);
-              pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\<rbrakk>
-             \<Longrightarrow>
-               path_image h1 \<subseteq> - cball z (e / 2) \<and>
-               path_image h2 \<subseteq> - cball z (e / 2) \<and>
-               (\<forall>f. f holomorphic_on - cball z (e / 2) \<longrightarrow> contour_integral h2 f = contour_integral h1 f)"
-    using contour_integral_nearby_ends [OF oc \<gamma> ppag] by metis
-  obtain p where p: "valid_path p" "z \<notin> path_image p"
-                    "pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma>"
-              and pg: "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (\<gamma> t - p t) < min d e / 2"
-              and pi: "contour_integral p (\<lambda>x. 1 / (x - z)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z"
-    using winding_number [OF \<gamma> z, of "min d e / 2"] \<open>d>0\<close> \<open>e>0\<close> by auto
-  { fix w
-    assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2"
-    then have wnotp: "w \<notin> path_image p"
-      using cbg \<open>d>0\<close> \<open>e>0\<close>
-      apply (simp add: path_image_def cball_def dist_norm, clarify)
-      apply (frule pg)
-      apply (drule_tac c="\<gamma> x" in subsetD)
-      apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l)
-      done
-    have wnotg: "w \<notin> path_image \<gamma>"
-      using cbg e2 \<open>e>0\<close> by (force simp: dist_norm norm_minus_commute)
-    { fix k::real
-      assume k: "k>0"
-      then obtain q where q: "valid_path q" "w \<notin> path_image q"
-                             "pathstart q = pathstart \<gamma> \<and> pathfinish q = pathfinish \<gamma>"
-                    and qg: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (\<gamma> t - q t) < min k (min d e) / 2"
-                    and qi: "contour_integral q (\<lambda>u. 1 / (u - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
-        using winding_number [OF \<gamma> wnotg, of "min k (min d e) / 2"] \<open>d>0\<close> \<open>e>0\<close> k
-        by (force simp: min_divide_distrib_right)
-      have "contour_integral p (\<lambda>u. 1 / (u - w)) = contour_integral q (\<lambda>u. 1 / (u - w))"
-        apply (rule pi_eq [OF \<open>valid_path q\<close> \<open>valid_path p\<close>, THEN conjunct2, THEN conjunct2, rule_format])
-        apply (frule pg)
-        apply (frule qg)
-        using p q \<open>d>0\<close> e2
-        apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
-        done
-      then have "contour_integral p (\<lambda>x. 1 / (x - w)) = complex_of_real (2 * pi) * \<i> * winding_number \<gamma> w"
-        by (simp add: pi qi)
-    } note pip = this
-    have "path p"
-      using p by (simp add: valid_path_imp_path)
-    then have "winding_number p w = winding_number \<gamma> w"
-      apply (rule winding_number_unique [OF _ wnotp])
-      apply (rule_tac x=p in exI)
-      apply (simp add: p wnotp min_divide_distrib_right pip)
-      done
-  } note wnwn = this
-  obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \<subseteq> - path_image p"
-    using p open_contains_cball [of "- path_image p"]
-    by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path])
-  obtain L
-    where "L>0"
-      and L: "\<And>f B. \<lbrakk>f holomorphic_on - cball z (3 / 4 * pe);
-                      \<forall>z \<in> - cball z (3 / 4 * pe). cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
-                      cmod (contour_integral p f) \<le> L * B"
-    using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \<open>valid_path p\<close> by force
-  { fix e::real and w::complex
-    assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)"
-    then have [simp]: "w \<notin> path_image p"
-      using cbp p(2) \<open>0 < pe\<close>
-      by (force simp: dist_norm norm_minus_commute path_image_def cball_def)
-    have [simp]: "contour_integral p (\<lambda>x. 1/(x - w)) - contour_integral p (\<lambda>x. 1/(x - z)) =
-                  contour_integral p (\<lambda>x. 1/(x - w) - 1/(x - z))"
-      by (simp add: p contour_integrable_inversediff contour_integral_diff)
-    { fix x
-      assume pe: "3/4 * pe < cmod (z - x)"
-      have "cmod (w - x) < pe/4 + cmod (z - x)"
-        by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1))
-      then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp
-      have "cmod (z - x) \<le> cmod (z - w) + cmod (w - x)"
-        using norm_diff_triangle_le by blast
-      also have "... < pe/4 + cmod (w - x)"
-        using w by (simp add: norm_minus_commute)
-      finally have "pe/2 < cmod (w - x)"
-        using pe by auto
-      then have "(pe/2)^2 < cmod (w - x) ^ 2"
-        apply (rule power_strict_mono)
-        using \<open>pe>0\<close> by auto
-      then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2"
-        by (simp add: power_divide)
-      have "8 * L * cmod (w - z) < e * pe\<^sup>2"
-        using w \<open>L>0\<close> by (simp add: field_simps)
-      also have "... < e * 4 * cmod (w - x) * cmod (w - x)"
-        using pe2 \<open>e>0\<close> by (simp add: power2_eq_square)
-      also have "... < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))"
-        using wx
-        apply (rule mult_strict_left_mono)
-        using pe2 e not_less_iff_gr_or_eq by fastforce
-      finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)"
-        by simp
-      also have "... \<le> e * cmod (w - x) * cmod (z - x)"
-         using e by simp
-      finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" .
-      have "L * cmod (1 / (x - w) - 1 / (x - z)) \<le> e"
-        apply (cases "x=z \<or> x=w")
-        using pe \<open>pe>0\<close> w \<open>L>0\<close>
-        apply (force simp: norm_minus_commute)
-        using wx w(2) \<open>L>0\<close> pe pe2 Lwz
-        apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square)
-        done
-    } note L_cmod_le = this
-    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)"
-      apply (rule L)
-      using \<open>pe>0\<close> w
-      apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros)
-      using \<open>pe>0\<close> w \<open>L>0\<close>
-      apply (auto simp: cball_def dist_norm field_simps L_cmod_le  simp del: less_divide_eq_numeral1 le_divide_eq_numeral1)
-      done
-    have "cmod (contour_integral p (\<lambda>x. 1 / (x - w)) - contour_integral p (\<lambda>x. 1 / (x - z))) < 2*e"
-      apply (simp add:)
-      apply (rule le_less_trans [OF *])
-      using \<open>L>0\<close> e
-      apply (force simp: field_simps)
-      done
-    then have "cmod (winding_number p w - winding_number p z) < e"
-      using pi_ge_two e
-      by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans)
-  } note cmod_wn_diff = this
-  then have "isCont (winding_number p) z"
-    apply (simp add: continuous_at_eps_delta, clarify)
-    apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI)
-    using \<open>pe>0\<close> \<open>L>0\<close>
-    apply (simp add: dist_norm cmod_wn_diff)
-    done
-  then show ?thesis
-    apply (rule continuous_transform_within [where d = "min d e / 2"])
-    apply (auto simp: \<open>d>0\<close> \<open>e>0\<close> dist_norm wnwn)
-    done
-qed
-
-corollary continuous_on_winding_number:
-    "path \<gamma> \<Longrightarrow> continuous_on (- path_image \<gamma>) (\<lambda>w. winding_number \<gamma> w)"
-  by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number)
-
-
-subsection\<open>The winding number is constant on a connected region\<close>
-
-lemma winding_number_constant:
-  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and cs: "connected s" and sg: "s \<inter> path_image \<gamma> = {}"
-    shows "\<exists>k. \<forall>z \<in> s. winding_number \<gamma> z = k"
-proof -
-  { fix y z
-    assume ne: "winding_number \<gamma> y \<noteq> winding_number \<gamma> z"
-    assume "y \<in> s" "z \<in> s"
-    then have "winding_number \<gamma> y \<in> \<int>"  "winding_number \<gamma> z \<in>  \<int>"
-      using integer_winding_number [OF \<gamma> loop] sg \<open>y \<in> s\<close> by auto
-    with ne have "1 \<le> cmod (winding_number \<gamma> y - winding_number \<gamma> z)"
-      by (auto simp: Ints_def of_int_diff [symmetric] simp del: of_int_diff)
-  } note * = this
-  show ?thesis
-    apply (rule continuous_discrete_range_constant [OF cs])
-    using continuous_on_winding_number [OF \<gamma>] sg
-    apply (metis Diff_Compl Diff_eq_empty_iff continuous_on_subset)
-    apply (rule_tac x=1 in exI)
-    apply (auto simp: *)
-    done
-qed
-
-lemma winding_number_eq:
-     "\<lbrakk>path \<gamma>; pathfinish \<gamma> = pathstart \<gamma>; w \<in> s; z \<in> s; connected s; s \<inter> path_image \<gamma> = {}\<rbrakk>
-      \<Longrightarrow> winding_number \<gamma> w = winding_number \<gamma> z"
-using winding_number_constant by fastforce
-
-lemma open_winding_number_levelsets:
-  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-    shows "open {z. z \<notin> path_image \<gamma> \<and> winding_number \<gamma> z = k}"
-proof -
-  have op: "open (- path_image \<gamma>)"
-    by (simp add: closed_path_image \<gamma> open_Compl)
-  { fix z assume z: "z \<notin> path_image \<gamma>" and k: "k = winding_number \<gamma> z"
-    obtain e where e: "e>0" "ball z e \<subseteq> - path_image \<gamma>"
-      using open_contains_ball [of "- path_image \<gamma>"] op z
-      by blast
-    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"
-      apply (rule_tac x=e in exI)
-      using e apply (simp add: dist_norm ball_def norm_minus_commute)
-      apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where s = "ball z e"])
-      done
-  } then
-  show ?thesis
-    by (auto simp: open_dist)
-qed
-
-subsection\<open>Winding number is zero "outside" a curve, in various senses\<close>
-
-lemma winding_number_zero_in_outside:
-  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and z: "z \<in> outside (path_image \<gamma>)"
-    shows "winding_number \<gamma> z = 0"
-proof -
-  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
-    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
-  obtain w::complex where w: "w \<notin> ball 0 (B + 1)"
-    by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real)
-  have "- ball 0 (B + 1) \<subseteq> outside (path_image \<gamma>)"
-    apply (rule outside_subset_convex)
-    using B subset_ball by auto
-  then have wout: "w \<in> outside (path_image \<gamma>)"
-    using w by blast
-  moreover obtain k where "\<And>z. z \<in> outside (path_image \<gamma>) \<Longrightarrow> winding_number \<gamma> z = k"
-    using winding_number_constant [OF \<gamma> loop, of "outside(path_image \<gamma>)"] connected_outside
-    by (metis DIM_complex bounded_path_image dual_order.refl \<gamma> outside_no_overlap)
-  ultimately have "winding_number \<gamma> z = winding_number \<gamma> w"
-    using z by blast
-  also have "... = 0"
-  proof -
-    have wnot: "w \<notin> path_image \<gamma>"  using wout by (simp add: outside_def)
-    { fix e::real assume "0<e"
-      obtain p where p: "polynomial_function p" "pathstart p = pathstart \<gamma>" "pathfinish p = pathfinish \<gamma>"
-                 and pg1: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < 1)"
-                 and pge: "(\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk> \<Longrightarrow> cmod (p t - \<gamma> t) < e)"
-        using path_approx_polynomial_function [OF \<gamma>, of "min 1 e"] \<open>e>0\<close> by force
-      have pip: "path_image p \<subseteq> ball 0 (B + 1)"
-        using B
-        apply (clarsimp simp add: path_image_def dist_norm ball_def)
-        apply (frule (1) pg1)
-        apply (fastforce dest: norm_add_less)
-        done
-      then have "w \<notin> path_image p"  using w by blast
-      then have "\<exists>p. valid_path p \<and> w \<notin> path_image p \<and>
-                     pathstart p = pathstart \<gamma> \<and> pathfinish p = pathfinish \<gamma> \<and>
-                     (\<forall>t\<in>{0..1}. cmod (\<gamma> t - p t) < e) \<and> contour_integral p (\<lambda>wa. 1 / (wa - w)) = 0"
-        apply (rule_tac x=p in exI)
-        apply (simp add: p valid_path_polynomial_function)
-        apply (intro conjI)
-        using pge apply (simp add: norm_minus_commute)
-        apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]])
-        apply (rule holomorphic_intros | simp add: dist_norm)+
-        using mem_ball_0 w apply blast
-        using p apply (simp_all add: valid_path_polynomial_function loop pip)
-        done
-    }
-    then show ?thesis
-      by (auto intro: winding_number_unique [OF \<gamma>] simp add: wnot)
-  qed
-  finally show ?thesis .
-qed
-
-lemma winding_number_zero_outside:
-    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; z \<notin> s; path_image \<gamma> \<subseteq> s\<rbrakk> \<Longrightarrow> winding_number \<gamma> z = 0"
-  by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside)
-
-lemma winding_number_zero_at_infinity:
-  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-    shows "\<exists>B. \<forall>z. B \<le> norm z \<longrightarrow> winding_number \<gamma> z = 0"
-proof -
-  obtain B::real where "0 < B" and B: "path_image \<gamma> \<subseteq> ball 0 B"
-    using bounded_subset_ballD [OF bounded_path_image [OF \<gamma>]] by auto
-  then show ?thesis
-    apply (rule_tac x="B+1" in exI, clarify)
-    apply (rule winding_number_zero_outside [OF \<gamma> convex_cball [of 0 B] loop])
-    apply (meson less_add_one mem_cball_0 not_le order_trans)
-    using ball_subset_cball by blast
-qed
-
-lemma winding_number_zero_point:
-    "\<lbrakk>path \<gamma>; convex s; pathfinish \<gamma> = pathstart \<gamma>; open s; path_image \<gamma> \<subseteq> s\<rbrakk>
-     \<Longrightarrow> \<exists>z. z \<in> s \<and> winding_number \<gamma> z = 0"
-  using outside_compact_in_open [of "path_image \<gamma>" s] path_image_nonempty winding_number_zero_in_outside
-  by (fastforce simp add: compact_path_image)
-
-
-text\<open>If a path winds round a set, it winds rounds its inside.\<close>
-lemma winding_number_around_inside:
-  assumes \<gamma>: "path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-      and cls: "closed s" and cos: "connected s" and s_disj: "s \<inter> path_image \<gamma> = {}"
-      and z: "z \<in> s" and wn_nz: "winding_number \<gamma> z \<noteq> 0" and w: "w \<in> s \<union> inside s"
-    shows "winding_number \<gamma> w = winding_number \<gamma> z"
-proof -
-  have ssb: "s \<subseteq> inside(path_image \<gamma>)"
-  proof
-    fix x :: complex
-    assume "x \<in> s"
-    hence "x \<notin> path_image \<gamma>"
-      by (meson disjoint_iff_not_equal s_disj)
-    thus "x \<in> inside (path_image \<gamma>)"
-      using \<open>x \<in> s\<close> 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)
-qed
-  show ?thesis
-    apply (rule winding_number_eq [OF \<gamma> loop w])
-    using z apply blast
-    apply (simp add: cls connected_with_inside cos)
-    apply (simp add: Int_Un_distrib2 s_disj, safe)
-    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)
- qed
-
-
-text\<open>Bounding a WN by 1/2 for a path and point in opposite halfspaces.\<close>
-lemma winding_number_subpath_continuous:
-  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
-    shows "continuous_on {0..1} (\<lambda>x. winding_number(subpath 0 x \<gamma>) z)"
-proof -
-  have *: "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
-         winding_number (subpath 0 x \<gamma>) z"
-         if x: "0 \<le> x" "x \<le> 1" for x
-  proof -
-    have "integral {0..x} (\<lambda>t. vector_derivative \<gamma> (at t) / (\<gamma> t - z)) / (2 * of_real pi * \<i>) =
-          1 / (2*pi*\<i>) * contour_integral (subpath 0 x \<gamma>) (\<lambda>w. 1/(w - z))"
-      using assms x
-      apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff])
-      done
-    also have "... = winding_number (subpath 0 x \<gamma>) z"
-      apply (subst winding_number_valid_path)
-      using assms x
-      apply (simp_all add: path_image_subpath valid_path_subpath)
-      by (force simp: path_image_def)
-    finally show ?thesis .
-  qed
-  show ?thesis
-    apply (rule continuous_on_eq
-                 [where f = "\<lambda>x. 1 / (2*pi*\<i>) *
-                                 integral {0..x} (\<lambda>t. 1/(\<gamma> t - z) * vector_derivative \<gamma> (at t))"])
-    apply (rule continuous_intros)+
-    apply (rule indefinite_integral_continuous)
-    apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on])
-      using assms
-    apply (simp add: *)
-    done
-qed
-
-lemma winding_number_ivt_pos:
-    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "0 \<le> w" "w \<le> Re(winding_number \<gamma> z)"
-      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
-  apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
-  apply (simp add:)
-  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
-  using assms
-  apply (auto simp: path_image_def image_def)
-  done
-
-lemma winding_number_ivt_neg:
-    assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>" and "Re(winding_number \<gamma> z) \<le> w" "w \<le> 0"
-      shows "\<exists>t \<in> {0..1}. Re(winding_number(subpath 0 t \<gamma>) z) = w"
-  apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right])
-  apply (simp add:)
-  apply (rule winding_number_subpath_continuous [OF \<gamma> z])
-  using assms
-  apply (auto simp: path_image_def image_def)
-  done
-
-lemma winding_number_ivt_abs:
-    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>"
-      shows "\<exists>t \<in> {0..1}. \<bar>Re (winding_number (subpath 0 t \<gamma>) z)\<bar> = w"
-  using assms winding_number_ivt_pos [of \<gamma> z w] winding_number_ivt_neg [of \<gamma> z "-w"]
-  by force
-
-lemma winding_number_lt_half_lemma:
-  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}"
-    shows "Re(winding_number \<gamma> z) < 1/2"
-proof -
-  { assume "Re(winding_number \<gamma> z) \<ge> 1/2"
-    then obtain t::real where t: "0 \<le> t" "t \<le> 1" and sub12: "Re (winding_number (subpath 0 t \<gamma>) z) = 1/2"
-      using winding_number_ivt_pos [OF \<gamma> z, of "1/2"] by auto
-    have gt: "\<gamma> t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \<gamma>) z)))) * (\<gamma> 0 - z))"
-      using winding_number_exp_2pi [of "subpath 0 t \<gamma>" z]
-      apply (simp add: t \<gamma> valid_path_imp_path)
-      using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12)
-    have "b < a \<bullet> \<gamma> 0"
-    proof -
-      have "\<gamma> 0 \<in> {c. b < a \<bullet> c}"
-        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one)
-      thus ?thesis
-        by blast
-    qed
-    moreover have "b < a \<bullet> \<gamma> t"
-    proof -
-      have "\<gamma> t \<in> {c. b < a \<bullet> c}"
-        by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t)
-      thus ?thesis
-        by blast
-    qed
-    ultimately have "0 < a \<bullet> (\<gamma> 0 - z)" "0 < a \<bullet> (\<gamma> t - z)" using az
-      by (simp add: inner_diff_right)+
-    then have False
-      by (simp add: gt inner_mult_right mult_less_0_iff)
-  }
-  then show ?thesis by force
-qed
-
-lemma winding_number_lt_half:
-  assumes "valid_path \<gamma>" "a \<bullet> z \<le> b" "path_image \<gamma> \<subseteq> {w. a \<bullet> w > b}"
-    shows "\<bar>Re (winding_number \<gamma> z)\<bar> < 1/2"
-proof -
-  have "z \<notin> path_image \<gamma>" using assms by auto
-  with assms show ?thesis
-    apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1)
-    apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \<gamma> z a b]
-                 winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath)
-    done
-qed
-
-lemma winding_number_le_half:
-  assumes \<gamma>: "valid_path \<gamma>" and z: "z \<notin> path_image \<gamma>"
-      and anz: "a \<noteq> 0" and azb: "a \<bullet> z \<le> b" and pag: "path_image \<gamma> \<subseteq> {w. a \<bullet> w \<ge> b}"
-    shows "\<bar>Re (winding_number \<gamma> z)\<bar> \<le> 1/2"
-proof -
-  { assume wnz_12: "\<bar>Re (winding_number \<gamma> z)\<bar> > 1/2"
-    have "isCont (winding_number \<gamma>) z"
-      by (metis continuous_at_winding_number valid_path_imp_path \<gamma> z)
-    then obtain d where "d>0" and d: "\<And>x'. dist x' z < d \<Longrightarrow> dist (winding_number \<gamma> x') (winding_number \<gamma> z) < \<bar>Re(winding_number \<gamma> z)\<bar> - 1/2"
-      using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast
-    define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a"
-    have *: "a \<bullet> z' \<le> b - d / 3 * cmod a"
-      unfolding z'_def inner_mult_right' divide_inverse
-      apply (simp add: divide_simps algebra_simps dot_square_norm power2_eq_square anz)
-      apply (metis \<open>0 < d\<close> add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral)
-      done
-    have "cmod (winding_number \<gamma> z' - winding_number \<gamma> z) < \<bar>Re (winding_number \<gamma> z)\<bar> - 1/2"
-      using d [of z'] anz \<open>d>0\<close> by (simp add: dist_norm z'_def)
-    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - cmod (winding_number \<gamma> z' - winding_number \<gamma> z)"
-      by simp
-    then have "1/2 < \<bar>Re (winding_number \<gamma> z)\<bar> - \<bar>Re (winding_number \<gamma> z') - Re (winding_number \<gamma> z)\<bar>"
-      using abs_Re_le_cmod [of "winding_number \<gamma> z' - winding_number \<gamma> z"] by simp
-    then have wnz_12': "\<bar>Re (winding_number \<gamma> z')\<bar> > 1/2"
-      by linarith
-    moreover have "\<bar>Re (winding_number \<gamma> z')\<bar> < 1/2"
-      apply (rule winding_number_lt_half [OF \<gamma> *])
-      using azb \<open>d>0\<close> pag
-      apply (auto simp: add_strict_increasing anz divide_simps algebra_simps dest!: subsetD)
-      done
-    ultimately have False
-      by simp
-  }
-  then show ?thesis by force
-qed
-
-lemma winding_number_lt_half_linepath: "z \<notin> closed_segment a b \<Longrightarrow> \<bar>Re (winding_number (linepath a b) z)\<bar> < 1/2"
-  using separating_hyperplane_closed_point [of "closed_segment a b" z]
-  apply auto
-  apply (simp add: closed_segment_def)
-  apply (drule less_imp_le)
-  apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]])
-  apply (auto simp: segment)
-  done
-
-
-text\<open> Positivity of WN for a linepath.\<close>
-lemma winding_number_linepath_pos_lt:
-    assumes "0 < Im ((b - a) * cnj (b - z))"
-      shows "0 < Re(winding_number(linepath a b) z)"
-proof -
-  have z: "z \<notin> path_image (linepath a b)"
-    using assms
-    by (simp add: closed_segment_def) (force simp: algebra_simps)
-  show ?thesis
-    apply (rule winding_number_pos_lt [OF valid_path_linepath z assms])
-    apply (simp add: linepath_def algebra_simps)
-    done
-qed
-
-
-subsection\<open>Cauchy's integral formula, again for a convex enclosing set.\<close>
-
-lemma Cauchy_integral_formula_weak:
-    assumes s: "convex s" and "finite k" and conf: "continuous_on s f"
-        and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
-        and z: "z \<in> interior s - k" and vpg: "valid_path \<gamma>"
-        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
-proof -
-  obtain f' where f': "(f has_field_derivative f') (at z)"
-    using fcd [OF z] by (auto simp: field_differentiable_def)
-  have pas: "path_image \<gamma> \<subseteq> s" and znotin: "z \<notin> path_image \<gamma>" using pasz by blast+
-  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
-  proof (cases "x = z")
-    case True then show ?thesis
-      apply (simp add: continuous_within)
-      apply (rule Lim_transform_away_within [of _ "z+1" _ "\<lambda>w::complex. (f w - f z)/(w - z)"])
-      using has_field_derivative_at_within DERIV_within_iff f'
-      apply (fastforce simp add:)+
-      done
-  next
-    case False
-    then have dxz: "dist x z > 0" by auto
-    have cf: "continuous (at x within s) f"
-      using conf continuous_on_eq_continuous_within that by blast
-    have "continuous (at x within s) (\<lambda>w. (f w - f z) / (w - z))"
-      by (rule cf continuous_intros | simp add: False)+
-    then show ?thesis
-      apply (rule continuous_transform_within [OF _ dxz that, of "\<lambda>w::complex. (f w - f z)/(w - z)"])
-      apply (force simp: dist_commute)
-      done
-  qed
-  have fink': "finite (insert z k)" using \<open>finite k\<close> by blast
-  have *: "((\<lambda>w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
-    apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop])
-    using c apply (force simp: continuous_on_eq_continuous_within)
-    apply (rename_tac w)
-    apply (rule_tac d="dist w z" and f = "\<lambda>w. (f w - f z)/(w - z)" in field_differentiable_transform_within)
-    apply (simp_all add: dist_pos_lt dist_commute)
-    apply (metis less_irrefl)
-    apply (rule derivative_intros fcd | simp)+
-    done
-  show ?thesis
-    apply (rule has_contour_integral_eq)
-    using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *]
-    apply (auto simp: mult_ac divide_simps)
-    done
-qed
-
-theorem Cauchy_integral_formula_convex_simple:
-    "\<lbrakk>convex s; f holomorphic_on s; z \<in> interior s; valid_path \<gamma>; path_image \<gamma> \<subseteq> s - {z};
-      pathfinish \<gamma> = pathstart \<gamma>\<rbrakk>
-     \<Longrightarrow> ((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
-  apply (rule Cauchy_integral_formula_weak [where k = "{}"])
-  using holomorphic_on_imp_continuous_on
-  by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE)
-
-
-subsection\<open>Homotopy forms of Cauchy's theorem\<close>
-
-proposition Cauchy_theorem_homotopic:
-    assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h"
-        and "open s" and f: "f holomorphic_on s"
-        and vpg: "valid_path g" and vph: "valid_path h"
-    shows "contour_integral g f = contour_integral h f"
-proof -
-  have pathsf: "linked_paths atends g h"
-    using hom  by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop)
-  obtain k :: "real \<times> real \<Rightarrow> complex"
-    where contk: "continuous_on ({0..1} \<times> {0..1}) k"
-      and ks: "k ` ({0..1} \<times> {0..1}) \<subseteq> s"
-      and k [simp]: "\<forall>x. k (0, x) = g x" "\<forall>x. k (1, x) = h x"
-      and ksf: "\<forall>t\<in>{0..1}. linked_paths atends g (\<lambda>x. k (t, x))"
-      using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm)
-  have ucontk: "uniformly_continuous_on ({0..1} \<times> {0..1}) k"
-    by (blast intro: compact_Times compact_uniformly_continuous [OF contk])
-  { fix t::real assume t: "t \<in> {0..1}"
-    have pak: "path (k o (\<lambda>u. (t, u)))"
-      unfolding path_def
-      apply (rule continuous_intros continuous_on_subset [OF contk])+
-      using t by force
-    have pik: "path_image (k \<circ> Pair t) \<subseteq> s"
-      using ks t by (auto simp: path_image_def)
-    obtain e where "e>0" and e:
-         "\<And>g h. \<lbrakk>valid_path g; valid_path h;
-                  \<forall>u\<in>{0..1}. cmod (g u - (k \<circ> Pair t) u) < e \<and> cmod (h u - (k \<circ> Pair t) u) < e;
-                  linked_paths atends g h\<rbrakk>
-                 \<Longrightarrow> contour_integral h f = contour_integral g f"
-      using contour_integral_nearby [OF \<open>open s\<close> pak pik, of atends] f by metis
-    obtain d where "d>0" and d:
-        "\<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"
-      by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \<open>e>0\<close>)
-    { fix t1 t2
-      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"
-      have no2: "\<And>g1 k1 kt. \<lbrakk>norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\<rbrakk> \<Longrightarrow> norm(g1 - kt) < e"
-        using \<open>e > 0\<close>
-        apply (rule_tac y = k1 in norm_triangle_half_l)
-        apply (auto simp: norm_minus_commute intro: order_less_trans)
-        done
-      have "\<exists>d>0. \<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
-                          (\<forall>u\<in>{0..1}. cmod (g1 u - k (t1, u)) < d \<and> cmod (g2 u - k (t2, u)) < d) \<and>
-                          linked_paths atends g1 g2 \<longrightarrow>
-                          contour_integral g2 f = contour_integral g1 f"
-        apply (rule_tac x="e/4" in exI)
-        using t t1 t2 ltd \<open>e > 0\<close>
-        apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1)
-        done
-    }
-    then have "\<exists>e. 0 < e \<and>
-              (\<forall>t1 t2. t1 \<in> {0..1} \<and> t2 \<in> {0..1} \<and> \<bar>t1 - t\<bar> < e \<and> \<bar>t2 - t\<bar> < e
-                \<longrightarrow> (\<exists>d. 0 < d \<and>
-                     (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
-                       (\<forall>u \<in> {0..1}.
-                          norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
-                          linked_paths atends g1 g2
-                          \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
-      by (rule_tac x=d in exI) (simp add: \<open>d > 0\<close>)
-  }
-  then obtain ee where ee:
-       "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0 \<and>
-          (\<forall>t1 t2. t1 \<in> {0..1} \<longrightarrow> t2 \<in> {0..1} \<longrightarrow> \<bar>t1 - t\<bar> < ee t \<longrightarrow> \<bar>t2 - t\<bar> < ee t
-            \<longrightarrow> (\<exists>d. 0 < d \<and>
-                 (\<forall>g1 g2. valid_path g1 \<and> valid_path g2 \<and>
-                   (\<forall>u \<in> {0..1}.
-                      norm(g1 u - k((t1,u))) < d \<and> norm(g2 u - k((t2,u))) < d) \<and>
-                      linked_paths atends g1 g2
-                      \<longrightarrow> contour_integral g2 f = contour_integral g1 f)))"
-    by metis
-  note ee_rule = ee [THEN conjunct2, rule_format]
-  define C where "C = (\<lambda>t. ball t (ee t / 3)) ` {0..1}"
-  have "\<forall>t \<in> C. open t" by (simp add: C_def)
-  moreover have "{0..1} \<subseteq> \<Union>C"
-    using ee [THEN conjunct1] by (auto simp: C_def dist_norm)
-  ultimately obtain C' where C': "C' \<subseteq> C" "finite C'" and C'01: "{0..1} \<subseteq> \<Union>C'"
-    by (rule compactE [OF compact_interval])
-  define kk where "kk = {t \<in> {0..1}. ball t (ee t / 3) \<in> C'}"
-  have kk01: "kk \<subseteq> {0..1}" by (auto simp: kk_def)
-  define e where "e = Min (ee ` kk)"
-  have C'_eq: "C' = (\<lambda>t. ball t (ee t / 3)) ` kk"
-    using C' by (auto simp: kk_def C_def)
-  have ee_pos[simp]: "\<And>t. t \<in> {0..1} \<Longrightarrow> ee t > 0"
-    by (simp add: kk_def ee)
-  moreover have "finite kk"
-    using \<open>finite C'\<close> kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD)
-  moreover have "kk \<noteq> {}" using \<open>{0..1} \<subseteq> \<Union>C'\<close> C'_eq by force
-  ultimately have "e > 0"
-    using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def)
-  then obtain N::nat where "N > 0" and N: "1/N < e/3"
-    by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral)
-  have e_le_ee: "\<And>i. i \<in> kk \<Longrightarrow> e \<le> ee i"
-    using \<open>finite kk\<close> by (simp add: e_def Min_le_iff [of "ee ` kk"])
-  have plus: "\<exists>t \<in> kk. x \<in> ball t (ee t / 3)" if "x \<in> {0..1}" for x
-    using C' subsetD [OF C'01 that]  unfolding C'_eq by blast
-  have [OF order_refl]:
-      "\<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
-                        \<longrightarrow> contour_integral j f = contour_integral g f)"
-       if "n \<le> N" for n
-  using that
-  proof (induct n)
-    case 0 show ?case using ee_rule [of 0 0 0]
-      apply clarsimp
-      apply (rule_tac x=d in exI, safe)
-      by (metis diff_self vpg norm_zero)
-  next
-    case (Suc n)
-    then have N01: "n/N \<in> {0..1}" "(Suc n)/N \<in> {0..1}"  by auto
-    then obtain t where t: "t \<in> kk" "n/N \<in> ball t (ee t / 3)"
-      using plus [of "n/N"] by blast
-    then have nN_less: "\<bar>n/N - t\<bar> < ee t"
-      by (simp add: dist_norm del: less_divide_eq_numeral1)
-    have n'N_less: "\<bar>real (Suc n) / real N - t\<bar> < ee t"
-      using t N \<open>N > 0\<close> e_le_ee [of t]
-      by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps)
-    have t01: "t \<in> {0..1}" using \<open>kk \<subseteq> {0..1}\<close> \<open>t \<in> kk\<close> by blast
-    obtain d1 where "d1 > 0" and d1:
-        "\<And>g1 g2. \<lbrakk>valid_path g1; valid_path g2;
-                   \<forall>u\<in>{0..1}. cmod (g1 u - k (n/N, u)) < d1 \<and> cmod (g2 u - k ((Suc n) / N, u)) < d1;
-                   linked_paths atends g1 g2\<rbrakk>
-                   \<Longrightarrow> contour_integral g2 f = contour_integral g1 f"
-      using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce
-    have "n \<le> N" using Suc.prems by auto
-    with Suc.hyps
-    obtain d2 where "d2 > 0"
-      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>
-                     \<Longrightarrow> contour_integral j f = contour_integral g f"
-        by auto
-    have "continuous_on {0..1} (k o (\<lambda>u. (n/N, u)))"
-      apply (rule continuous_intros continuous_on_subset [OF contk])+
-      using N01 by auto
-    then have pkn: "path (\<lambda>u. k (n/N, u))"
-      by (simp add: path_def)
-    have min12: "min d1 d2 > 0" by (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>)
-    obtain p where "polynomial_function p"
-        and psf: "pathstart p = pathstart (\<lambda>u. k (n/N, u))"
-                 "pathfinish p = pathfinish (\<lambda>u. k (n/N, u))"
-        and pk_le:  "\<And>t. t\<in>{0..1} \<Longrightarrow> cmod (p t - k (n/N, t)) < min d1 d2"
-      using path_approx_polynomial_function [OF pkn min12] by blast
-    then have vpp: "valid_path p" using valid_path_polynomial_function by blast
-    have lpa: "linked_paths atends g p"
-      by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf)
-    show ?case
-      apply (rule_tac x="min d1 d2" in exI)
-      apply (simp add: \<open>0 < d1\<close> \<open>0 < d2\<close>, clarify)
-      apply (rule_tac s="contour_integral p f" in trans)
-      using pk_le N01(1) ksf pathfinish_def pathstart_def
-      apply (force intro!: vpp d1 simp add: linked_paths_def psf ksf)
-      using pk_le N01 apply (force intro!: vpp d2 lpa simp add: linked_paths_def psf ksf)
-      done
-  qed
-  then obtain d where "0 < d"
-                       "\<And>j. valid_path j \<and> (\<forall>u \<in> {0..1}. norm(j u - k (1,u)) < d) \<and>
-                            linked_paths atends g j
-                            \<Longrightarrow> contour_integral j f = contour_integral g f"
-    using \<open>N>0\<close> by auto
-  then have "linked_paths atends g h \<Longrightarrow> contour_integral h f = contour_integral g f"
-    using \<open>N>0\<close> vph by fastforce
-  then show ?thesis
-    by (simp add: pathsf)
-qed
-
-proposition Cauchy_theorem_homotopic_paths:
-    assumes hom: "homotopic_paths s g h"
-        and "open s" and f: "f holomorphic_on s"
-        and vpg: "valid_path g" and vph: "valid_path h"
-    shows "contour_integral g f = contour_integral h f"
-  using Cauchy_theorem_homotopic [of True s g h] assms by simp
-
-proposition Cauchy_theorem_homotopic_loops:
-    assumes hom: "homotopic_loops s g h"
-        and "open s" and f: "f holomorphic_on s"
-        and vpg: "valid_path g" and vph: "valid_path h"
-    shows "contour_integral g f = contour_integral h f"
-  using Cauchy_theorem_homotopic [of False s g h] assms by simp
-
-lemma has_contour_integral_newpath:
-    "\<lbrakk>(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\<rbrakk>
-     \<Longrightarrow> (f has_contour_integral y) g"
-  using has_contour_integral_integral contour_integral_unique by auto
-
-lemma Cauchy_theorem_null_homotopic:
-     "\<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"
-  apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp)
-  using contour_integrable_holomorphic_simple
-    apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset)
-  by (simp add: Cauchy_theorem_homotopic_loops)
-
-
-
-subsection\<open>More winding number properties\<close>
-
-text\<open>including the fact that it's +-1 inside a simple closed curve.\<close>
-
-lemma winding_number_homotopic_paths:
-    assumes "homotopic_paths (-{z}) g h"
-      shows "winding_number g z = winding_number h z"
-proof -
-  have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto
-  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
-    using homotopic_paths_imp_subset [OF assms] by auto
-  ultimately obtain d e where "d > 0" "e > 0"
-      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>
-            \<Longrightarrow> homotopic_paths (-{z}) g p"
-      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>
-            \<Longrightarrow> homotopic_paths (-{z}) h q"
-    using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force
-  obtain p where p:
-       "valid_path p" "z \<notin> path_image p"
-       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
-       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
-       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
-    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
-  obtain q where q:
-       "valid_path q" "z \<notin> path_image q"
-       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
-       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
-       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
-    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
-  have gp: "homotopic_paths (- {z}) g p"
-    by (simp add: d p valid_path_imp_path norm_minus_commute gp_less)
-  have hq: "homotopic_paths (- {z}) h q"
-    by (simp add: e q valid_path_imp_path norm_minus_commute hq_less)
-  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
-    apply (rule Cauchy_theorem_homotopic_paths [of "-{z}"])
-    apply (blast intro: homotopic_paths_trans homotopic_paths_sym gp hq assms)
-    apply (auto intro!: holomorphic_intros simp: p q)
-    done
-  then show ?thesis
-    by (simp add: pap paq)
-qed
-
-lemma winding_number_homotopic_loops:
-    assumes "homotopic_loops (-{z}) g h"
-      shows "winding_number g z = winding_number h z"
-proof -
-  have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto
-  moreover have pag: "z \<notin> path_image g" and pah: "z \<notin> path_image h"
-    using homotopic_loops_imp_subset [OF assms] by auto
-  moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h"
-    using homotopic_loops_imp_loop [OF assms] by auto
-  ultimately obtain d e where "d > 0" "e > 0"
-      and d: "\<And>p. \<lbrakk>path p; pathfinish p = pathstart p; \<forall>t\<in>{0..1}. norm (p t - g t) < d\<rbrakk>
-            \<Longrightarrow> homotopic_loops (-{z}) g p"
-      and e: "\<And>q. \<lbrakk>path q; pathfinish q = pathstart q; \<forall>t\<in>{0..1}. norm (q t - h t) < e\<rbrakk>
-            \<Longrightarrow> homotopic_loops (-{z}) h q"
-    using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force
-  obtain p where p:
-       "valid_path p" "z \<notin> path_image p"
-       "pathstart p = pathstart g" "pathfinish p = pathfinish g"
-       and gp_less:"\<forall>t\<in>{0..1}. cmod (g t - p t) < d"
-       and pap: "contour_integral p (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number g z"
-    using winding_number [OF \<open>path g\<close> pag \<open>0 < d\<close>] by blast
-  obtain q where q:
-       "valid_path q" "z \<notin> path_image q"
-       "pathstart q = pathstart h" "pathfinish q = pathfinish h"
-       and hq_less: "\<forall>t\<in>{0..1}. cmod (h t - q t) < e"
-       and paq:  "contour_integral q (\<lambda>w. 1 / (w - z)) = complex_of_real (2 * pi) * \<i> * winding_number h z"
-    using winding_number [OF \<open>path h\<close> pah \<open>0 < e\<close>] by blast
-  have gp: "homotopic_loops (- {z}) g p"
-    by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path)
-  have hq: "homotopic_loops (- {z}) h q"
-    by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path)
-  have "contour_integral p (\<lambda>w. 1/(w - z)) = contour_integral q (\<lambda>w. 1/(w - z))"
-    apply (rule Cauchy_theorem_homotopic_loops [of "-{z}"])
-    apply (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms)
-    apply (auto intro!: holomorphic_intros simp: p q)
-    done
-  then show ?thesis
-    by (simp add: pap paq)
-qed
-
-lemma winding_number_paths_linear_eq:
-  "\<lbrakk>path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g;
-    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
-        \<Longrightarrow> winding_number h z = winding_number g z"
-  by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths elim: )
-
-lemma winding_number_loops_linear_eq:
-  "\<lbrakk>path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h;
-    \<And>t. t \<in> {0..1} \<Longrightarrow> z \<notin> closed_segment (g t) (h t)\<rbrakk>
-        \<Longrightarrow> winding_number h z = winding_number g z"
-  by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops elim: )
-
-lemma winding_number_nearby_paths_eq:
-     "\<lbrakk>path g; path h;
-      pathstart h = pathstart g; pathfinish h = pathfinish g;
-      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
-      \<Longrightarrow> winding_number h z = winding_number g z"
-  by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq)
-
-lemma winding_number_nearby_loops_eq:
-     "\<lbrakk>path g; path h;
-      pathfinish g = pathstart g;
-        pathfinish h = pathstart h;
-      \<And>t. t \<in> {0..1} \<Longrightarrow> norm(h t - g t) < norm(g t - z)\<rbrakk>
-      \<Longrightarrow> winding_number h z = winding_number g z"
-  by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq)
-
-
-proposition winding_number_subpath_combine:
-    "\<lbrakk>path g; z \<notin> path_image g;
-      u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk>
-      \<Longrightarrow> winding_number (subpath u v g) z + winding_number (subpath v w g) z =
-          winding_number (subpath u w g) z"
-apply (rule trans [OF winding_number_join [THEN sym]
-                      winding_number_homotopic_paths [OF homotopic_join_subpaths]])
-apply (auto dest: path_image_subpath_subset)
-done
-
-
-subsection\<open>Partial circle path\<close>
-
-definition part_circlepath :: "[complex, real, real, real, real] \<Rightarrow> complex"
-  where "part_circlepath z r s t \<equiv> \<lambda>x. z + of_real r * exp (\<i> * of_real (linepath s t x))"
-
-lemma pathstart_part_circlepath [simp]:
-     "pathstart(part_circlepath z r s t) = z + r*exp(\<i> * s)"
-by (metis part_circlepath_def pathstart_def pathstart_linepath)
-
-lemma pathfinish_part_circlepath [simp]:
-     "pathfinish(part_circlepath z r s t) = z + r*exp(\<i>*t)"
-by (metis part_circlepath_def pathfinish_def pathfinish_linepath)
-
-proposition has_vector_derivative_part_circlepath [derivative_intros]:
-    "((part_circlepath z r s t) has_vector_derivative
-      (\<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)))
-     (at x within X)"
-  apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real)
-  apply (rule has_vector_derivative_real_complex)
-  apply (rule derivative_eq_intros | simp)+
-  done
-
-corollary vector_derivative_part_circlepath:
-    "vector_derivative (part_circlepath z r s t) (at x) =
-       \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
-  using has_vector_derivative_part_circlepath vector_derivative_at by blast
-
-corollary vector_derivative_part_circlepath01:
-    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
-     \<Longrightarrow> vector_derivative (part_circlepath z r s t) (at x within {0..1}) =
-          \<i> * r * (of_real t - of_real s) * exp(\<i> * linepath s t x)"
-  using has_vector_derivative_part_circlepath
-  by (auto simp: vector_derivative_at_within_ivl)
-
-lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)"
-  apply (simp add: valid_path_def)
-  apply (rule C1_differentiable_imp_piecewise)
-  apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath
-              intro!: continuous_intros)
-  done
-
-lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)"
-  by (simp add: valid_path_imp_path)
-
-proposition path_image_part_circlepath:
-  assumes "s \<le> t"
-    shows "path_image (part_circlepath z r s t) = {z + r * exp(\<i> * of_real x) | x. s \<le> x \<and> x \<le> t}"
-proof -
-  { fix z::real
-    assume "0 \<le> z" "z \<le> 1"
-    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"
-      apply (rule_tac x="(1 - z) * s + z * t" in exI)
-      apply (simp add: linepath_def scaleR_conv_of_real algebra_simps)
-      apply (rule conjI)
-      using mult_right_mono apply blast
-      using affine_ineq  by (metis "mult.commute")
-  }
-  moreover
-  { fix z
-    assume "s \<le> z" "z \<le> t"
-    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}"
-      apply (rule_tac x="(z - s)/(t - s)" in image_eqI)
-      apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq)
-      apply (auto simp: algebra_simps divide_simps)
-      done
-  }
-  ultimately show ?thesis
-    by (fastforce simp add: path_image_def part_circlepath_def)
-qed
-
-corollary path_image_part_circlepath_subset:
-    "\<lbrakk>s \<le> t; 0 \<le> r\<rbrakk> \<Longrightarrow> path_image(part_circlepath z r s t) \<subseteq> sphere z r"
-by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult)
-
-proposition in_path_image_part_circlepath:
-  assumes "w \<in> path_image(part_circlepath z r s t)" "s \<le> t" "0 \<le> r"
-    shows "norm(w - z) = r"
-proof -
-  have "w \<in> {c. dist z c = r}"
-    by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms)
-  thus ?thesis
-    by (simp add: dist_norm norm_minus_commute)
-qed
-
-proposition finite_bounded_log: "finite {z::complex. norm z \<le> b \<and> exp z = w}"
-proof (cases "w = 0")
-  case True then show ?thesis by auto
-next
-  case False
-  have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \<i>) \<le> b + cmod (Ln w)}"
-    apply (simp add: norm_mult finite_int_iff_bounded_le)
-    apply (rule_tac x="\<lfloor>(b + cmod (Ln w)) / (2*pi)\<rfloor>" in exI)
-    apply (auto simp: divide_simps le_floor_iff)
-    done
-  have [simp]: "\<And>P f. {z. P z \<and> (\<exists>n. z = f n)} = f ` {n. P (f n)}"
-    by blast
-  show ?thesis
-    apply (subst exp_Ln [OF False, symmetric])
-    apply (simp add: exp_eq)
-    using norm_add_leD apply (fastforce intro: finite_subset [OF _ *])
-    done
-qed
-
-lemma finite_bounded_log2:
-  fixes a::complex
-    assumes "a \<noteq> 0"
-    shows "finite {z. norm z \<le> b \<and> exp(a*z) = w}"
-proof -
-  have *: "finite ((\<lambda>z. z / a) ` {z. cmod z \<le> b * cmod a \<and> exp z = w})"
-    by (rule finite_imageI [OF finite_bounded_log])
-  show ?thesis
-    by (rule finite_subset [OF _ *]) (force simp: assms norm_mult)
-qed
-
-proposition has_contour_integral_bound_part_circlepath_strong:
-  assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)"
-      and "finite k" and le: "0 \<le> B" "0 < r" "s \<le> t"
-      and B: "\<And>x. x \<in> path_image(part_circlepath z r s t) - k \<Longrightarrow> norm(f x) \<le> B"
-    shows "cmod i \<le> B * r * (t - s)"
-proof -
-  consider "s = t" | "s < t" using \<open>s \<le> t\<close> by linarith
-  then show ?thesis
-  proof cases
-    case 1 with fi [unfolded has_contour_integral]
-    have "i = 0"  by (simp add: vector_derivative_part_circlepath)
-    with assms show ?thesis by simp
-  next
-    case 2
-    have [simp]: "\<bar>r\<bar> = r" using \<open>r > 0\<close> by linarith
-    have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s"
-      by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff)
-    have "finite (part_circlepath z r s t -` {y} \<inter> {0..1})" if "y \<in> k" for y
-    proof -
-      define w where "w = (y - z)/of_real r / exp(\<i> * of_real s)"
-      have fin: "finite (of_real -` {z. cmod z \<le> 1 \<and> exp (\<i> * complex_of_real (t - s) * z) = w})"
-        apply (rule finite_vimageI [OF finite_bounded_log2])
-        using \<open>s < t\<close> apply (auto simp: inj_of_real)
-        done
-      show ?thesis
-        apply (simp add: part_circlepath_def linepath_def vimage_def)
-        apply (rule finite_subset [OF _ fin])
-        using le
-        apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff)
-        done
-    qed
-    then have fin01: "finite ((part_circlepath z r s t) -` k \<inter> {0..1})"
-      by (rule finite_finite_vimage_IntI [OF \<open>finite k\<close>])
-    have **: "((\<lambda>x. if (part_circlepath z r s t x) \<in> k then 0
-                    else f(part_circlepath z r s t x) *
-                       vector_derivative (part_circlepath z r s t) (at x)) has_integral i)  {0..1}"
-      apply (rule has_integral_spike
-              [where f = "\<lambda>x. f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)"])
-      apply (rule negligible_finite [OF fin01])
-      using fi has_contour_integral
-      apply auto
-      done
-    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"
-      by (auto intro!: B [unfolded path_image_def image_def, simplified])
-    show ?thesis
-      apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified])
-      using assms apply force
-      apply (simp add: norm_mult vector_derivative_part_circlepath)
-      using le * "2" \<open>r > 0\<close> by auto
-  qed
-qed
-
-corollary has_contour_integral_bound_part_circlepath:
-      "\<lbrakk>(f has_contour_integral i) (part_circlepath z r s t);
-        0 \<le> B; 0 < r; s \<le> t;
-        \<And>x. x \<in> path_image(part_circlepath z r s t) \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
-       \<Longrightarrow> norm i \<le> B*r*(t - s)"
-  by (auto intro: has_contour_integral_bound_part_circlepath_strong)
-
-proposition contour_integrable_continuous_part_circlepath:
-     "continuous_on (path_image (part_circlepath z r s t)) f
-      \<Longrightarrow> f contour_integrable_on (part_circlepath z r s t)"
-  apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def)
-  apply (rule integrable_continuous_real)
-  apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl])
-  done
-
-proposition winding_number_part_circlepath_pos_less:
-  assumes "s < t" and no: "norm(w - z) < r"
-    shows "0 < Re (winding_number(part_circlepath z r s t) w)"
-proof -
-  have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2)
-  note valid_path_part_circlepath
-  moreover have " w \<notin> path_image (part_circlepath z r s t)"
-    using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def)
-  moreover have "0 < r * (t - s) * (r - cmod (w - z))"
-    using assms by (metis \<open>0 < r\<close> diff_gt_0_iff_gt mult_pos_pos)
-  ultimately show ?thesis
-    apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"])
-    apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac)
-    apply (rule mult_left_mono)+
-    using Re_Im_le_cmod [of "w-z" "linepath s t x" for x]
-    apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square])
-    using assms \<open>0 < r\<close> by auto
-qed
-
-proposition simple_path_part_circlepath:
-    "simple_path(part_circlepath z r s t) \<longleftrightarrow> (r \<noteq> 0 \<and> s \<noteq> t \<and> \<bar>s - t\<bar> \<le> 2*pi)"
-proof (cases "r = 0 \<or> s = t")
-  case True
-  then show ?thesis
-    apply (rule disjE)
-    apply (force simp: part_circlepath_def simple_path_def intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+
-    done
-next
-  case False then have "r \<noteq> 0" "s \<noteq> t" by auto
-  have *: "\<And>x y z s t. \<i>*((1 - x) * s + x * t) = \<i>*(((1 - y) * s + y * t)) + z  \<longleftrightarrow> \<i>*(x - y) * (t - s) = z"
-    by (simp add: algebra_simps)
-  have abs01: "\<And>x y::real. 0 \<le> x \<and> x \<le> 1 \<and> 0 \<le> y \<and> y \<le> 1
-                      \<Longrightarrow> (x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0 \<longleftrightarrow> \<bar>x - y\<bar> \<in> {0,1})"
-    by auto
-  have abs_away: "\<And>P. (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. P \<bar>x - y\<bar>) \<longleftrightarrow> (\<forall>x::real. 0 \<le> x \<and> x \<le> 1 \<longrightarrow> P x)"
-    by force
-  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>
-                  (\<exists>n. \<bar>x - y\<bar> * (t - s) = 2 * (of_int n * pi))"
-    by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real]
-                    intro: exI [where x = "-n" for n])
-  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"
-    apply (rule ccontr)
-    apply (drule_tac x="2*pi / \<bar>t - s\<bar>" in spec)
-    using False
-    apply (simp add: abs_minus_commute divide_simps)
-    apply (frule_tac x=1 in spec)
-    apply (drule_tac x="-1" in spec)
-    apply (simp add:)
-    done
-  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
-  proof -
-    have "t-s = 2 * (real_of_int n * pi)/x"
-      using that by (simp add: field_simps)
-    then show ?thesis by (metis abs_minus_commute)
-  qed
-  show ?thesis using False
-    apply (simp add: simple_path_def path_part_circlepath)
-    apply (simp add: part_circlepath_def linepath_def exp_eq  * ** abs01  del: Set.insert_iff)
-    apply (subst abs_away)
-    apply (auto simp: 1)
-    apply (rule ccontr)
-    apply (auto simp: 2 divide_simps abs_mult dest: of_int_leD)
-    done
-qed
-
-proposition arc_part_circlepath:
-  assumes "r \<noteq> 0" "s \<noteq> t" "\<bar>s - t\<bar> < 2*pi"
-    shows "arc (part_circlepath z r s t)"
-proof -
-  have *: "x = y" if eq: "\<i> * (linepath s t x) = \<i> * (linepath s t y) + 2 * of_int n * complex_of_real pi * \<i>"
-                  and x: "x \<in> {0..1}" and y: "y \<in> {0..1}" for x y n
-    proof -
-      have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi"
-        by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_divide_cancel_left eq)
-      then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))"
-        by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re])
-      then have st: "x \<noteq> y \<Longrightarrow> (s-t) = (of_int n * (pi * 2) / (y-x))"
-        by (force simp: field_simps)
-      show ?thesis
-        apply (rule ccontr)
-        using assms x y
-        apply (simp add: st abs_mult field_simps)
-        using st
-        apply (auto simp: dest: of_int_lessD)
-        done
-    qed
-  show ?thesis
-    using assms
-    apply (simp add: arc_def)
-    apply (simp add: part_circlepath_def inj_on_def exp_eq)
-    apply (blast intro: *)
-    done
-qed
-
-
-subsection\<open>Special case of one complete circle\<close>
-
-definition circlepath :: "[complex, real, real] \<Rightarrow> complex"
-  where "circlepath z r \<equiv> part_circlepath z r 0 (2*pi)"
-
-lemma circlepath: "circlepath z r = (\<lambda>x. z + r * exp(2 * of_real pi * \<i> * of_real x))"
-  by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps)
-
-lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r"
-  by (simp add: circlepath_def)
-
-lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r"
-  by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute)
-
-lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)"
-proof -
-  have "z + of_real r * exp (2 * pi * \<i> * (x + 1 / 2)) =
-        z + of_real r * exp (2 * pi * \<i> * x + pi * \<i>)"
-    by (simp add: divide_simps) (simp add: algebra_simps)
-  also have "... = z - r * exp (2 * pi * \<i> * x)"
-    by (simp add: exp_add)
-  finally show ?thesis
-    by (simp add: circlepath path_image_def sphere_def dist_norm)
-qed
-
-lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x"
-  using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x]
-  by (simp add: add.commute)
-
-lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)"
-  using circlepath_add1 [of z r "x-1/2"]
-  by (simp add: add.commute)
-
-lemma path_image_circlepath_minus_subset:
-     "path_image (circlepath z (-r)) \<subseteq> path_image (circlepath z r)"
-  apply (simp add: path_image_def image_def circlepath_minus, clarify)
-  apply (case_tac "xa \<le> 1/2", force)
-  apply (force simp add: circlepath_add_half)+
-  done
-
-lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)"
-  using path_image_circlepath_minus_subset by fastforce
-
-proposition has_vector_derivative_circlepath [derivative_intros]:
- "((circlepath z r) has_vector_derivative (2 * pi * \<i> * r * exp (2 * of_real pi * \<i> * of_real x)))
-   (at x within X)"
-  apply (simp add: circlepath_def scaleR_conv_of_real)
-  apply (rule derivative_eq_intros)
-  apply (simp add: algebra_simps)
-  done
-
-corollary vector_derivative_circlepath:
-   "vector_derivative (circlepath z r) (at x) =
-    2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
-using has_vector_derivative_circlepath vector_derivative_at by blast
-
-corollary vector_derivative_circlepath01:
-    "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk>
-     \<Longrightarrow> vector_derivative (circlepath z r) (at x within {0..1}) =
-          2 * pi * \<i> * r * exp(2 * of_real pi * \<i> * x)"
-  using has_vector_derivative_circlepath
-  by (auto simp: vector_derivative_at_within_ivl)
-
-lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)"
-  by (simp add: circlepath_def)
-
-lemma path_circlepath [simp]: "path (circlepath z r)"
-  by (simp add: valid_path_imp_path)
-
-lemma path_image_circlepath_nonneg:
-  assumes "0 \<le> r" shows "path_image (circlepath z r) = sphere z r"
-proof -
-  have *: "x \<in> (\<lambda>u. z + (cmod (x - z)) * exp (\<i> * (of_real u * (of_real pi * 2)))) ` {0..1}" for x
-  proof (cases "x = z")
-    case True then show ?thesis by force
-  next
-    case False
-    define w where "w = x - z"
-    then have "w \<noteq> 0" by (simp add: False)
-    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)"
-      using cis_conv_exp complex_eq_iff by auto
-    show ?thesis
-      apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"])
-      apply (simp add: divide_simps \<open>w \<noteq> 0\<close> cmod_power2 [symmetric])
-      apply (rule_tac x="t / (2*pi)" in image_eqI)
-      apply (simp add: divide_simps \<open>w \<noteq> 0\<close>)
-      using False **
-      apply (auto simp: w_def)
-      done
-  qed
-  show ?thesis
-    unfolding circlepath path_image_def sphere_def dist_norm
-    by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *)
-qed
-
-proposition path_image_circlepath [simp]:
-    "path_image (circlepath z r) = sphere z \<bar>r\<bar>"
-  using path_image_circlepath_minus
-  by (force simp add: path_image_circlepath_nonneg abs_if)
-
-lemma has_contour_integral_bound_circlepath_strong:
-      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
-        finite k; 0 \<le> B; 0 < r;
-        \<And>x. \<lbrakk>norm(x - z) = r; x \<notin> k\<rbrakk> \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
-        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
-  unfolding circlepath_def
-  by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong)
-
-corollary has_contour_integral_bound_circlepath:
-      "\<lbrakk>(f has_contour_integral i) (circlepath z r);
-        0 \<le> B; 0 < r; \<And>x. norm(x - z) = r \<Longrightarrow> norm(f x) \<le> B\<rbrakk>
-        \<Longrightarrow> norm i \<le> B*(2*pi*r)"
-  by (auto intro: has_contour_integral_bound_circlepath_strong)
-
-proposition contour_integrable_continuous_circlepath:
-    "continuous_on (path_image (circlepath z r)) f
-     \<Longrightarrow> f contour_integrable_on (circlepath z r)"
-  by (simp add: circlepath_def contour_integrable_continuous_part_circlepath)
-
-lemma simple_path_circlepath: "simple_path(circlepath z r) \<longleftrightarrow> (r \<noteq> 0)"
-  by (simp add: circlepath_def simple_path_part_circlepath)
-
-lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \<Longrightarrow> w \<notin> path_image (circlepath z r)"
-  by (simp add: sphere_def dist_norm norm_minus_commute)
-
-proposition contour_integral_circlepath:
-     "0 < r \<Longrightarrow> contour_integral (circlepath z r) (\<lambda>w. 1 / (w - z)) = 2 * complex_of_real pi * \<i>"
-  apply (rule contour_integral_unique)
-  apply (simp add: has_contour_integral_def)
-  apply (subst has_integral_cong)
-  apply (simp add: vector_derivative_circlepath01)
-  using has_integral_const_real [of _ 0 1]
-  apply (force simp: circlepath)
-  done
-
-lemma winding_number_circlepath_centre: "0 < r \<Longrightarrow> winding_number (circlepath z r) z = 1"
-  apply (rule winding_number_unique_loop)
-  apply (simp_all add: sphere_def valid_path_imp_path)
-  apply (rule_tac x="circlepath z r" in exI)
-  apply (simp add: sphere_def contour_integral_circlepath)
-  done
-
-proposition winding_number_circlepath:
-  assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1"
-proof (cases "w = z")
-  case True then show ?thesis
-    using assms winding_number_circlepath_centre by auto
-next
-  case False
-  have [simp]: "r > 0"
-    using assms le_less_trans norm_ge_zero by blast
-  define r' where "r' = norm(w - z)"
-  have "r' < r"
-    by (simp add: assms r'_def)
-  have disjo: "cball z r' \<inter> sphere z r = {}"
-    using \<open>r' < r\<close> by (force simp: cball_def sphere_def)
-  have "winding_number(circlepath z r) w = winding_number(circlepath z r) z"
-    apply (rule winding_number_around_inside [where s = "cball z r'"])
-    apply (simp_all add: disjo order.strict_implies_order winding_number_circlepath_centre)
-    apply (simp_all add: False r'_def dist_norm norm_minus_commute)
-    done
-  also have "... = 1"
-    by (simp add: winding_number_circlepath_centre)
-  finally show ?thesis .
-qed
-
-
-text\<open> Hence the Cauchy formula for points inside a circle.\<close>
-
-theorem Cauchy_integral_circlepath:
-  assumes "continuous_on (cball z r) f" "f holomorphic_on (ball z r)" "norm(w - z) < r"
-  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
-         (circlepath z r)"
-proof -
-  have "r > 0"
-    using assms le_less_trans norm_ge_zero by blast
-  have "((\<lambda>u. f u / (u - w)) has_contour_integral (2 * pi) * \<i> * winding_number (circlepath z r) w * f w)
-        (circlepath z r)"
-    apply (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"])
-    using assms  \<open>r > 0\<close>
-    apply (simp_all add: dist_norm norm_minus_commute)
-    apply (metis at_within_interior dist_norm holomorphic_on_def interior_ball mem_ball norm_minus_commute)
-    apply (simp add: cball_def sphere_def dist_norm, clarify)
-    apply (simp add:)
-    by (metis dist_commute dist_norm less_irrefl)
-  then show ?thesis
-    by (simp add: winding_number_circlepath assms)
-qed
-
-corollary Cauchy_integral_circlepath_simple:
-  assumes "f holomorphic_on cball z r" "norm(w - z) < r"
-  shows "((\<lambda>u. f u/(u - w)) has_contour_integral (2 * of_real pi * \<i> * f w))
-         (circlepath z r)"
-using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath)
-
-
-lemma no_bounded_connected_component_imp_winding_number_zero:
-  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
-      and nb: "\<And>z. bounded (connected_component_set (- s) z) \<longrightarrow> z \<in> s"
-  shows "winding_number g z = 0"
-apply (rule winding_number_zero_in_outside)
-apply (simp_all add: assms)
-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)
-
-lemma no_bounded_path_component_imp_winding_number_zero:
-  assumes g: "path g" "path_image g \<subseteq> s" "pathfinish g = pathstart g" "z \<notin> s"
-      and nb: "\<And>z. bounded (path_component_set (- s) z) \<longrightarrow> z \<in> s"
-  shows "winding_number g z = 0"
-apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g])
-by (simp add: bounded_subset nb path_component_subset_connected_component)
-
-
-subsection\<open> Uniform convergence of path integral\<close>
-
-text\<open>Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\<close>
-
-proposition contour_integral_uniform_limit:
-  assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on \<gamma>) F"
-      and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image \<gamma>. norm(f n x - l x) < e) F"
-      and noleB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
-      and \<gamma>: "valid_path \<gamma>"
-      and [simp]: "~ (trivial_limit F)"
-  shows "l contour_integrable_on \<gamma>" "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
-proof -
-  have "0 \<le> B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one)
-  { fix e::real
-    assume "0 < e"
-    then have eB: "0 < e / (\<bar>B\<bar> + 1)" by simp
-    obtain a where fga: "\<And>x. x \<in> {0..1} \<Longrightarrow> cmod (f a (\<gamma> x) - l (\<gamma> x)) < e / (\<bar>B\<bar> + 1)"
-               and inta: "(\<lambda>t. f a (\<gamma> t) * vector_derivative \<gamma> (at t)) integrable_on {0..1}"
-      using eventually_happens [OF eventually_conj [OF ev_no [OF eB] ev_fint]]
-      by (fastforce simp: contour_integrable_on path_image_def)
-    have Ble: "B * e / (\<bar>B\<bar> + 1) \<le> e"
-      using \<open>0 \<le> B\<close>  \<open>0 < e\<close> by (simp add: divide_simps)
-    have "\<exists>h. (\<forall>x\<in>{0..1}. cmod (l (\<gamma> x) * vector_derivative \<gamma> (at x) - h x) \<le> e) \<and> h integrable_on {0..1}"
-      apply (rule_tac x="\<lambda>x. f (a::'a) (\<gamma> x) * vector_derivative \<gamma> (at x)" in exI)
-      apply (intro inta conjI ballI)
-      apply (rule order_trans [OF _ Ble])
-      apply (frule noleB)
-      apply (frule fga)
-      using \<open>0 \<le> B\<close>  \<open>0 < e\<close>
-      apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps)
-      apply (drule (1) mult_mono [OF less_imp_le])
-      apply (simp_all add: mult_ac)
-      done
-  }
-  then show lintg: "l contour_integrable_on \<gamma>"
-    apply (simp add: contour_integrable_on)
-    apply (blast intro: integrable_uniform_limit_real)
-    done
-  { fix e::real
-    define B' where "B' = B + 1"
-    have B': "B' > 0" "B' > B" using  \<open>0 \<le> B\<close> by (auto simp: B'_def)
-    assume "0 < e"
-    then have ev_no': "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image \<gamma>. 2 * cmod (f n x - l x) < e / B'"
-      using ev_no [of "e / B' / 2"] B' by (simp add: field_simps)
-    have ie: "integral {0..1::real} (\<lambda>x. e / 2) < e" using \<open>0 < e\<close> by simp
-    have *: "cmod (f x (\<gamma> t) * vector_derivative \<gamma> (at t) - l (\<gamma> t) * vector_derivative \<gamma> (at t)) \<le> e / 2"
-             if t: "t\<in>{0..1}" and leB': "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) < e / B'" for x t
-    proof -
-      have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) \<le> e * (B/ B')"
-        using mult_mono [OF less_imp_le [OF leB'] noleB] B' \<open>0 < e\<close> t by auto
-      also have "... < e"
-        by (simp add: B' \<open>0 < e\<close> mult_imp_div_pos_less)
-      finally have "2 * cmod (f x (\<gamma> t) - l (\<gamma> t)) * cmod (vector_derivative \<gamma> (at t)) < e" .
-      then show ?thesis
-        by (simp add: left_diff_distrib [symmetric] norm_mult)
-    qed
-    have "\<forall>\<^sub>F x in F. dist (contour_integral \<gamma> (f x)) (contour_integral \<gamma> l) < e"
-      apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]])
-      apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral)
-      apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric], clarify)
-      apply (rule le_less_trans [OF integral_norm_bound_integral ie])
-      apply (simp add: lintg integrable_diff contour_integrable_on [symmetric])
-      apply (blast intro: *)+
-      done
-  }
-  then show "((\<lambda>n. contour_integral \<gamma> (f n)) \<longlongrightarrow> contour_integral \<gamma> l) F"
-    by (rule tendstoI)
-qed
-
-proposition contour_integral_uniform_limit_circlepath:
-  assumes ev_fint: "eventually (\<lambda>n::'a. (f n) contour_integrable_on (circlepath z r)) F"
-      and ev_no: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> path_image (circlepath z r). norm(f n x - l x) < e) F"
-      and [simp]: "~ (trivial_limit F)" "0 < r"
-  shows "l contour_integrable_on (circlepath z r)" "((\<lambda>n. contour_integral (circlepath z r) (f n)) \<longlongrightarrow> contour_integral (circlepath z r) l) F"
-by (auto simp: vector_derivative_circlepath norm_mult intro: contour_integral_uniform_limit assms)
-
-
-subsection\<open> General stepping result for derivative formulas.\<close>
-
-lemma sum_sqs_eq:
-  fixes x::"'a::idom" shows "x * x + y * y = x * (y * 2) \<Longrightarrow> y = x"
-  by algebra
-
-proposition Cauchy_next_derivative:
-  assumes "continuous_on (path_image \<gamma>) f'"
-      and leB: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (vector_derivative \<gamma> (at t)) \<le> B"
-      and int: "\<And>w. w \<in> s - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f' u / (u - w)^k) has_contour_integral f w) \<gamma>"
-      and k: "k \<noteq> 0"
-      and "open s"
-      and \<gamma>: "valid_path \<gamma>"
-      and w: "w \<in> s - path_image \<gamma>"
-    shows "(\<lambda>u. f' u / (u - w)^(Suc k)) contour_integrable_on \<gamma>"
-      and "(f has_field_derivative (k * contour_integral \<gamma> (\<lambda>u. f' u/(u - w)^(Suc k))))
-           (at w)"  (is "?thes2")
-proof -
-  have "open (s - path_image \<gamma>)" using \<open>open s\<close> closed_valid_path_image \<gamma> by blast
-  then obtain d where "d>0" and d: "ball w d \<subseteq> s - path_image \<gamma>" using w
-    using open_contains_ball by blast
-  have [simp]: "\<And>n. cmod (1 + of_nat n) = 1 + of_nat n"
-    by (metis norm_of_nat of_nat_Suc)
-  have 1: "\<forall>\<^sub>F n in at w. (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k)
-                         contour_integrable_on \<gamma>"
-    apply (simp add: eventually_at)
-    apply (rule_tac x=d in exI)
-    apply (simp add: \<open>d > 0\<close> dist_norm field_simps, clarify)
-    apply (rule contour_integrable_div [OF contour_integrable_diff])
-    using int w d
-    apply (force simp:  dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+
-    done
-  have bim_g: "bounded (image f' (path_image \<gamma>))"
-    by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms)
-  then obtain C where "C > 0" and C: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cmod (f' (\<gamma> x)) \<le> C"
-    by (force simp: bounded_pos path_image_def)
-  have twom: "\<forall>\<^sub>F n in at w.
-               \<forall>x\<in>path_image \<gamma>.
-                cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e"
-         if "0 < e" for e
-  proof -
-    have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k)   < e"
-            if x: "x \<in> path_image \<gamma>" and "u \<noteq> w" and uwd: "cmod (u - w) < d/2"
-                and uw_less: "cmod (u - w) < e * (d / 2) ^ (k+2) / (1 + real k)"
-            for u x
-    proof -
-      define ff where [abs_def]:
-        "ff n w =
-          (if n = 0 then inverse(x - w)^k
-           else if n = 1 then k / (x - w)^(Suc k)
-           else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w
-      have km1: "\<And>z::complex. z \<noteq> 0 \<Longrightarrow> z ^ (k - Suc 0) = z ^ k / z"
-        by (simp add: field_simps) (metis Suc_pred \<open>k \<noteq> 0\<close> neq0_conv power_Suc)
-      have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d / 2))"
-              if "z \<in> ball w (d / 2)" "i \<le> 1" for i z
-      proof -
-        have "z \<notin> path_image \<gamma>"
-          using \<open>x \<in> path_image \<gamma>\<close> d that ball_divide_subset_numeral by blast
-        then have xz[simp]: "x \<noteq> z" using \<open>x \<in> path_image \<gamma>\<close> by blast
-        then have neq: "x * x + z * z \<noteq> x * (z * 2)"
-          by (blast intro: dest!: sum_sqs_eq)
-        with xz have "\<And>v. v \<noteq> 0 \<Longrightarrow> (x * x + z * z) * v \<noteq> (x * (z * 2) * v)" by auto
-        then have neqq: "\<And>v. v \<noteq> 0 \<Longrightarrow> x * (x * v) + z * (z * v) \<noteq> x * (z * (2 * v))"
-          by (simp add: algebra_simps)
-        show ?thesis using \<open>i \<le> 1\<close>
-          apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe)
-          apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+
-          done
-      qed
-      { fix a::real and b::real assume ab: "a > 0" "b > 0"
-        then have "k * (1 + real k) * (1 / a) \<le> k * (1 + real k) * (4 / b) \<longleftrightarrow> b \<le> 4 * a"
-          apply (subst mult_le_cancel_left_pos)
-          using \<open>k \<noteq> 0\<close>
-          apply (auto simp: divide_simps)
-          done
-        with ab have "real k * (1 + real k) / a \<le> (real k * 4 + real k * real k * 4) / b \<longleftrightarrow> b \<le> 4 * a"
-          by (simp add: field_simps)
-      } note canc = this
-      have ff2: "cmod (ff (Suc 1) v) \<le> real (k * (k + 1)) / (d / 2) ^ (k + 2)"
-                if "v \<in> ball w (d / 2)" for v
-      proof -
-        have "d/2 \<le> cmod (x - v)" using d x that
-          apply (simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps, clarify)
-          apply (drule subsetD)
-           prefer 2 apply blast
-          apply (metis norm_minus_commute norm_triangle_half_r CollectI)
-          done
-        then have "d \<le> cmod (x - v) * 2"
-          by (simp add: divide_simps)
-        then have dpow_le: "d ^ (k+2) \<le> (cmod (x - v) * 2) ^ (k+2)"
-          using \<open>0 < d\<close> order_less_imp_le power_mono by blast
-        have "x \<noteq> v" using that
-          using \<open>x \<in> path_image \<gamma>\<close> ball_divide_subset_numeral d by fastforce
-        then show ?thesis
-        using \<open>d > 0\<close>
-        apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc)
-        using dpow_le
-        apply (simp add: algebra_simps divide_simps mult_less_0_iff)
-        done
-      qed
-      have ub: "u \<in> ball w (d / 2)"
-        using uwd by (simp add: dist_commute dist_norm)
-      have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
-                  \<le> (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d / 2) ^ k))"
-        using complex_taylor [OF _ ff1 ff2 _ ub, of w, simplified]
-        by (simp add: ff_def \<open>0 < d\<close>)
-      then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
-                  \<le> (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
-        by (simp add: field_simps)
-      then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k)))
-                 / (cmod (u - w) * real k)
-                  \<le> (1 + real k) * cmod (u - w) / (d / 2) ^ (k+2)"
-        using \<open>k \<noteq> 0\<close> \<open>u \<noteq> w\<close> by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq)
-      also have "... < e"
-        using uw_less \<open>0 < d\<close> by (simp add: mult_ac divide_simps)
-      finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k)))
-                        / cmod ((u - w) * real k)   <   e"
-        by (simp add: norm_mult)
-      have "x \<noteq> u"
-        using uwd \<open>0 < d\<close> x d by (force simp: dist_norm ball_def norm_minus_commute)
-      show ?thesis
-        apply (rule le_less_trans [OF _ e])
-        using \<open>k \<noteq> 0\<close> \<open>x \<noteq> u\<close>  \<open>u \<noteq> w\<close>
-        apply (simp add: field_simps norm_divide [symmetric])
-        done
-    qed
-    show ?thesis
-      unfolding eventually_at
-      apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI)
-      apply (force simp: \<open>d > 0\<close> dist_norm that simp del: power_Suc intro: *)
-      done
-  qed
-  have 2: "\<forall>\<^sub>F n in at w.
-              \<forall>x\<in>path_image \<gamma>.
-               cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e"
-          if "0 < e" for e
-  proof -
-    have *: "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
-                        f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) < e"
-              if ec: "cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
-                      inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k) < e / C"
-                 and x: "0 \<le> x" "x \<le> 1"
-              for u x
-    proof (cases "(f' (\<gamma> x)) = 0")
-      case True then show ?thesis by (simp add: \<open>0 < e\<close>)
-    next
-      case False
-      have "cmod (f' (\<gamma> x) * (inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
-                        f' (\<gamma> x) / ((\<gamma> x - w) * (\<gamma> x - w) ^ k)) =
-            cmod (f' (\<gamma> x) * ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
-                             inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k))"
-        by (simp add: field_simps)
-      also have "... = cmod (f' (\<gamma> x)) *
-                       cmod ((inverse (\<gamma> x - u) ^ k - inverse (\<gamma> x - w) ^ k) / ((u - w) * k) -
-                             inverse (\<gamma> x - w) * inverse (\<gamma> x - w) ^ k)"
-        by (simp add: norm_mult)
-      also have "... < cmod (f' (\<gamma> x)) * (e/C)"
-        apply (rule mult_strict_left_mono [OF ec])
-        using False by simp
-      also have "... \<le> e" using C
-        by (metis False \<open>0 < e\<close> frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff)
-      finally show ?thesis .
-    qed
-    show ?thesis
-      using twom [OF divide_pos_pos [OF that \<open>C > 0\<close>]]   unfolding path_image_def
-      by (force intro: * elim: eventually_mono)
-  qed
-  show "(\<lambda>u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \<gamma>"
-    by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
-  have *: "(\<lambda>n. contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k))
-           \<midarrow>w\<rightarrow> contour_integral \<gamma> (\<lambda>u. f' u / (u - w) ^ (Suc k))"
-    by (rule contour_integral_uniform_limit [OF 1 2 leB \<gamma>]) auto
-  have **: "contour_integral \<gamma> (\<lambda>x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) =
-              (f u - f w) / (u - w) / k"
-           if "dist u w < d" for u
-    apply (rule contour_integral_unique)
-    apply (simp add: diff_divide_distrib algebra_simps)
-    apply (rule has_contour_integral_diff; rule has_contour_integral_div; simp add: field_simps; rule int)
-    apply (metis contra_subsetD d dist_commute mem_ball that)
-    apply (rule w)
-    done
-  show ?thes2
-    apply (simp add: DERIV_within_iff del: power_Suc)
-    apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \<open>0 < d\<close> ])
-    apply (simp add: \<open>k \<noteq> 0\<close> **)
-    done
-qed
-
-corollary Cauchy_next_derivative_circlepath:
-  assumes contf: "continuous_on (path_image (circlepath z r)) f"
-      and int: "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)"
-      and k: "k \<noteq> 0"
-      and w: "w \<in> ball z r"
-    shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
-           (is "?thes1")
-      and "(g has_field_derivative (k * contour_integral (circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)))) (at w)"
-           (is "?thes2")
-proof -
-  have "r > 0" using w
-    using ball_eq_empty by fastforce
-  have wim: "w \<in> ball z r - path_image (circlepath z r)"
-    using w by (auto simp: dist_norm)
-  show ?thes1 ?thes2
-    by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \<bar>r\<bar>"];
-        auto simp: vector_derivative_circlepath norm_mult)+
-qed
-
-
-text\<open> In particular, the first derivative formula.\<close>
-
-proposition Cauchy_derivative_integral_circlepath:
-  assumes contf: "continuous_on (cball z r) f"
-      and holf: "f holomorphic_on ball z r"
-      and w: "w \<in> ball z r"
-    shows "(\<lambda>u. f u/(u - w)^2) contour_integrable_on (circlepath z r)"
-           (is "?thes1")
-      and "(f has_field_derivative (1 / (2 * of_real pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u / (u - w)^2))) (at w)"
-           (is "?thes2")
-proof -
-  have [simp]: "r \<ge> 0" using w
-    using ball_eq_empty by fastforce
-  have f: "continuous_on (path_image (circlepath z r)) f"
-    by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def)
-  have int: "\<And>w. dist z w < r \<Longrightarrow>
-                 ((\<lambda>u. f u / (u - w)) has_contour_integral (\<lambda>x. 2 * of_real pi * \<i> * f x) w) (circlepath z r)"
-    by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute)
-  show ?thes1
-    apply (simp add: power2_eq_square)
-    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified])
-    apply (blast intro: int)
-    done
-  have "((\<lambda>x. 2 * of_real pi * \<i> * f x) has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2)) (at w)"
-    apply (simp add: power2_eq_square)
-    apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\<lambda>x. 2 * of_real pi * \<i> * f x", simplified])
-    apply (blast intro: int)
-    done
-  then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)^2) / (2 * of_real pi * \<i>)) (at w)"
-    by (rule DERIV_cdivide [where f = "\<lambda>x. 2 * of_real pi * \<i> * f x" and c = "2 * of_real pi * \<i>", simplified])
-  show ?thes2
-    by simp (rule fder)
-qed
-
-subsection\<open> Existence of all higher derivatives.\<close>
-
-proposition derivative_is_holomorphic:
-  assumes "open s"
-      and fder: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z)"
-    shows "f' holomorphic_on s"
-proof -
-  have *: "\<exists>h. (f' has_field_derivative h) (at z)" if "z \<in> s" for z
-  proof -
-    obtain r where "r > 0" and r: "cball z r \<subseteq> s"
-      using open_contains_cball \<open>z \<in> s\<close> \<open>open s\<close> by blast
-    then have holf_cball: "f holomorphic_on cball z r"
-      apply (simp add: holomorphic_on_def)
-      using field_differentiable_at_within field_differentiable_def fder by blast
-    then have "continuous_on (path_image (circlepath z r)) f"
-      using \<open>r > 0\<close> by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on])
-    then have contfpi: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1/(2 * of_real pi*\<i>) * f x)"
-      by (auto intro: continuous_intros)+
-    have contf_cball: "continuous_on (cball z r) f" using holf_cball
-      by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset)
-    have holf_ball: "f holomorphic_on ball z r" using holf_cball
-      using ball_subset_cball holomorphic_on_subset by blast
-    { fix w  assume w: "w \<in> ball z r"
-      have intf: "(\<lambda>u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r"
-        by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
-      have fder': "(f has_field_derivative 1 / (2 * of_real pi * \<i>) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2))
-                  (at w)"
-        by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball])
-      have f'_eq: "f' w = contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>)"
-        using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder])
-      have "((\<lambda>u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \<i>)) has_contour_integral
-                contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
-                (circlepath z r)"
-        by (rule Cauchy_Integral_Thm.has_contour_integral_div [OF has_contour_integral_integral [OF intf]])
-      then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral
-                contour_integral (circlepath z r) (\<lambda>u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \<i>))
-                (circlepath z r)"
-        by (simp add: algebra_simps)
-      then have "((\<lambda>u. f u / (2 * of_real pi * \<i> * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)"
-        by (simp add: f'_eq)
-    } note * = this
-    show ?thesis
-      apply (rule exI)
-      apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified])
-      apply (simp_all add: \<open>0 < r\<close> * dist_norm)
-      done
-  qed
-  show ?thesis
-    by (simp add: holomorphic_on_open [OF \<open>open s\<close>] *)
-qed
-
-lemma holomorphic_deriv [holomorphic_intros]:
-    "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv f) holomorphic_on s"
-by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def)
-
-lemma analytic_deriv: "f analytic_on s \<Longrightarrow> (deriv f) analytic_on s"
-  using analytic_on_holomorphic holomorphic_deriv by auto
-
-lemma holomorphic_higher_deriv [holomorphic_intros]: "\<lbrakk>f holomorphic_on s; open s\<rbrakk> \<Longrightarrow> (deriv ^^ n) f holomorphic_on s"
-  by (induction n) (auto simp: holomorphic_deriv)
-
-lemma analytic_higher_deriv: "f analytic_on s \<Longrightarrow> (deriv ^^ n) f analytic_on s"
-  unfolding analytic_on_def using holomorphic_higher_deriv by blast
-
-lemma has_field_derivative_higher_deriv:
-     "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk>
-      \<Longrightarrow> ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)"
-by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply
-         funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def)
-
-lemma valid_path_compose_holomorphic:
-  assumes "valid_path g" and holo:"f holomorphic_on s" and "open s" "path_image g \<subseteq> s"
-  shows "valid_path (f o g)"
-proof (rule valid_path_compose[OF \<open>valid_path g\<close>])
-  fix x assume "x \<in> path_image g"
-  then show "\<exists>f'. (f has_field_derivative f') (at x)"
-    using holo holomorphic_on_open[OF \<open>open s\<close>] \<open>path_image g \<subseteq> s\<close> by auto
-next
-  have "deriv f holomorphic_on s"
-    using holomorphic_deriv holo \<open>open s\<close> by auto
-  then show "continuous_on (path_image g) (deriv f)"
-    using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto
-qed
-
-
-subsection\<open> Morera's theorem.\<close>
-
-lemma Morera_local_triangle_ball:
-  assumes "\<And>z. z \<in> s
-          \<Longrightarrow> \<exists>e a. 0 < e \<and> z \<in> ball a e \<and> continuous_on (ball a e) f \<and>
-                    (\<forall>b c. closed_segment b c \<subseteq> ball a e
-                           \<longrightarrow> contour_integral (linepath a b) f +
-                               contour_integral (linepath b c) f +
-                               contour_integral (linepath c a) f = 0)"
-  shows "f analytic_on s"
-proof -
-  { fix z  assume "z \<in> s"
-    with assms obtain e a where
-            "0 < e" and z: "z \<in> ball a e" and contf: "continuous_on (ball a e) f"
-        and 0: "\<And>b c. closed_segment b c \<subseteq> ball a e
-                      \<Longrightarrow> contour_integral (linepath a b) f +
-                          contour_integral (linepath b c) f +
-                          contour_integral (linepath c a) f = 0"
-      by fastforce
-    have az: "dist a z < e" using mem_ball z by blast
-    have sb_ball: "ball z (e - dist a z) \<subseteq> ball a e"
-      by (simp add: dist_commute ball_subset_ball_iff)
-    have "\<exists>e>0. f holomorphic_on ball z e"
-      apply (rule_tac x="e - dist a z" in exI)
-      apply (simp add: az)
-      apply (rule holomorphic_on_subset [OF _ sb_ball])
-      apply (rule derivative_is_holomorphic[OF open_ball])
-      apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a])
-         apply (simp_all add: 0 \<open>0 < e\<close>)
-      apply (meson \<open>0 < e\<close> centre_in_ball convex_ball convex_contains_segment mem_ball)
-      done
-  }
-  then show ?thesis
-    by (simp add: analytic_on_def)
-qed
-
-lemma Morera_local_triangle:
-  assumes "\<And>z. z \<in> s
-          \<Longrightarrow> \<exists>t. open t \<and> z \<in> t \<and> continuous_on t f \<and>
-                  (\<forall>a b c. convex hull {a,b,c} \<subseteq> t
-                              \<longrightarrow> contour_integral (linepath a b) f +
-                                  contour_integral (linepath b c) f +
-                                  contour_integral (linepath c a) f = 0)"
-  shows "f analytic_on s"
-proof -
-  { fix z  assume "z \<in> s"
-    with assms obtain t where
-            "open t" and z: "z \<in> t" and contf: "continuous_on t f"
-        and 0: "\<And>a b c. convex hull {a,b,c} \<subseteq> t
-                      \<Longrightarrow> contour_integral (linepath a b) f +
-                          contour_integral (linepath b c) f +
-                          contour_integral (linepath c a) f = 0"
-      by force
-    then obtain e where "e>0" and e: "ball z e \<subseteq> t"
-      using open_contains_ball by blast
-    have [simp]: "continuous_on (ball z e) f" using contf
-      using continuous_on_subset e by blast
-    have "\<exists>e a. 0 < e \<and>
-               z \<in> ball a e \<and>
-               continuous_on (ball a e) f \<and>
-               (\<forall>b c. closed_segment b c \<subseteq> ball a e \<longrightarrow>
-                      contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)"
-      apply (rule_tac x=e in exI)
-      apply (rule_tac x=z in exI)
-      apply (simp add: \<open>e > 0\<close>, clarify)
-      apply (rule 0)
-      apply (meson z \<open>0 < e\<close> centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset)
-      done
-  }
-  then show ?thesis
-    by (simp add: Morera_local_triangle_ball)
-qed
-
-proposition Morera_triangle:
-    "\<lbrakk>continuous_on s f; open s;
-      \<And>a b c. convex hull {a,b,c} \<subseteq> s
-              \<longrightarrow> contour_integral (linepath a b) f +
-                  contour_integral (linepath b c) f +
-                  contour_integral (linepath c a) f = 0\<rbrakk>
-     \<Longrightarrow> f analytic_on s"
-  using Morera_local_triangle by blast
-
-
-
-subsection\<open> Combining theorems for higher derivatives including Leibniz rule.\<close>
-
-lemma higher_deriv_linear [simp]:
-    "(deriv ^^ n) (\<lambda>w. c*w) = (\<lambda>z. if n = 0 then c*z else if n = 1 then c else 0)"
-  by (induction n) (auto simp: deriv_const deriv_linear)
-
-lemma higher_deriv_const [simp]: "(deriv ^^ n) (\<lambda>w. c) = (\<lambda>w. if n=0 then c else 0)"
-  by (induction n) (auto simp: deriv_const)
-
-lemma higher_deriv_ident [simp]:
-     "(deriv ^^ n) (\<lambda>w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)"
-  apply (induction n, simp)
-  apply (metis higher_deriv_linear lambda_one)
-  done
-
-corollary higher_deriv_id [simp]:
-     "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)"
-  by (simp add: id_def)
-
-lemma has_complex_derivative_funpow_1:
-     "\<lbrakk>(f has_field_derivative 1) (at z); f z = z\<rbrakk> \<Longrightarrow> (f^^n has_field_derivative 1) (at z)"
-  apply (induction n)
-  apply auto
-  apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one)
-  by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral)
-
-proposition higher_deriv_uminus:
-  assumes "f holomorphic_on s" "open s" and z: "z \<in> s"
-    shows "(deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
-using z
-proof (induction n arbitrary: z)
-  case 0 then show ?case by simp
-next
-  case (Suc n z)
-  have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
-    using Suc.prems assms has_field_derivative_higher_deriv by auto
-  show ?case
-    apply simp
-    apply (rule DERIV_imp_deriv)
-    apply (rule DERIV_transform_within_open [of "\<lambda>w. -((deriv ^^ n) f w)"])
-    apply (rule derivative_eq_intros | rule * refl assms Suc)+
-    apply (simp add: Suc)
-    done
-qed
-
-proposition higher_deriv_add:
-  fixes z::complex
-  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
-    shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
-using z
-proof (induction n arbitrary: z)
-  case 0 then show ?case by simp
-next
-  case (Suc n z)
-  have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
-          "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
-    using Suc.prems assms has_field_derivative_higher_deriv by auto
-  show ?case
-    apply simp
-    apply (rule DERIV_imp_deriv)
-    apply (rule DERIV_transform_within_open [of "\<lambda>w. (deriv ^^ n) f w + (deriv ^^ n) g w"])
-    apply (rule derivative_eq_intros | rule * refl assms Suc)+
-    apply (simp add: Suc)
-    done
-qed
-
-corollary higher_deriv_diff:
-  fixes z::complex
-  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
-    shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
-  apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add)
-  apply (subst higher_deriv_add)
-  using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus)
-  done
-
-
-lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))"
-  by (cases k) simp_all
-
-proposition higher_deriv_mult:
-  fixes z::complex
-  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
-    shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
-           (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
-using z
-proof (induction n arbitrary: z)
-  case 0 then show ?case by simp
-next
-  case (Suc n z)
-  have *: "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)"
-          "\<And>n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)"
-    using Suc.prems assms has_field_derivative_higher_deriv by auto
-  have sumeq: "(\<Sum>i = 0..n.
-               of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) =
-            g z * deriv ((deriv ^^ n) f) z + (\<Sum>i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))"
-    apply (simp add: bb algebra_simps setsum.distrib)
-    apply (subst (4) setsum_Suc_reindex)
-    apply (auto simp: algebra_simps Suc_diff_le intro: setsum.cong)
-    done
-  show ?case
-    apply (simp only: funpow.simps o_apply)
-    apply (rule DERIV_imp_deriv)
-    apply (rule DERIV_transform_within_open
-             [of "\<lambda>w. (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"])
-    apply (simp add: algebra_simps)
-    apply (rule DERIV_cong [OF DERIV_setsum])
-    apply (rule DERIV_cmult)
-    apply (auto simp: intro: DERIV_mult * sumeq \<open>open s\<close> Suc.prems Suc.IH [symmetric])
-    done
-qed
-
-
-proposition higher_deriv_transform_within_open:
-  fixes z::complex
-  assumes "f holomorphic_on s" "g holomorphic_on s" "open s" and z: "z \<in> s"
-      and fg: "\<And>w. w \<in> s \<Longrightarrow> f w = g w"
-    shows "(deriv ^^ i) f z = (deriv ^^ i) g z"
-using z
-by (induction i arbitrary: z)
-   (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms)
-
-proposition higher_deriv_compose_linear:
-  fixes z::complex
-  assumes f: "f holomorphic_on t" and s: "open s" and t: "open t" and z: "z \<in> s"
-      and fg: "\<And>w. w \<in> s \<Longrightarrow> u * w \<in> t"
-    shows "(deriv ^^ n) (\<lambda>w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)"
-using z
-proof (induction n arbitrary: z)
-  case 0 then show ?case by simp
-next
-  case (Suc n z)
-  have holo0: "f holomorphic_on op * u ` s"
-    by (meson fg f holomorphic_on_subset image_subset_iff)
-  have holo1: "(\<lambda>w. f (u * w)) holomorphic_on s"
-    apply (rule holomorphic_on_compose [where g=f, unfolded o_def])
-    apply (rule holo0 holomorphic_intros)+
-    done
-  have holo2: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on s"
-    apply (rule holomorphic_intros)+
-    apply (rule holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def])
-    apply (rule holomorphic_intros)
-    apply (rule holomorphic_on_subset [where s=t])
-    apply (rule holomorphic_intros assms)+
-    apply (blast intro: fg)
-    done
-  have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z)) z"
-    apply (rule complex_derivative_transform_within_open [OF _ holo2 s Suc.prems])
-    apply (rule holomorphic_higher_deriv [OF holo1 s])
-    apply (simp add: Suc.IH)
-    done
-  also have "... = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z)) z"
-    apply (rule deriv_cmult)
-    apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems])
-    apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def])
-    apply (simp add: analytic_on_linear)
-    apply (simp add: analytic_on_open f holomorphic_higher_deriv t)
-    apply (blast intro: fg)
-    done
-  also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)"
-      apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "op*u", unfolded o_def])
-      apply (rule derivative_intros)
-      using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv t apply blast
-      apply (simp add: deriv_linear)
-      done
-  finally show ?case
-    by simp
-qed
-
-lemma higher_deriv_add_at:
-  assumes "f analytic_on {z}" "g analytic_on {z}"
-    shows "(deriv ^^ n) (\<lambda>w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z"
-proof -
-  have "f analytic_on {z} \<and> g analytic_on {z}"
-    using assms by blast
-  with higher_deriv_add show ?thesis
-    by (auto simp: analytic_at_two)
-qed
-
-lemma higher_deriv_diff_at:
-  assumes "f analytic_on {z}" "g analytic_on {z}"
-    shows "(deriv ^^ n) (\<lambda>w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z"
-proof -
-  have "f analytic_on {z} \<and> g analytic_on {z}"
-    using assms by blast
-  with higher_deriv_diff show ?thesis
-    by (auto simp: analytic_at_two)
-qed
-
-lemma higher_deriv_uminus_at:
-   "f analytic_on {z}  \<Longrightarrow> (deriv ^^ n) (\<lambda>w. -(f w)) z = - ((deriv ^^ n) f z)"
-  using higher_deriv_uminus
-    by (auto simp: analytic_at)
-
-lemma higher_deriv_mult_at:
-  assumes "f analytic_on {z}" "g analytic_on {z}"
-    shows "(deriv ^^ n) (\<lambda>w. f w * g w) z =
-           (\<Sum>i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)"
-proof -
-  have "f analytic_on {z} \<and> g analytic_on {z}"
-    using assms by blast
-  with higher_deriv_mult show ?thesis
-    by (auto simp: analytic_at_two)
-qed
-
-
-text\<open> Nonexistence of isolated singularities and a stronger integral formula.\<close>
-
-proposition no_isolated_singularity:
-  fixes z::complex
-  assumes f: "continuous_on s f" and holf: "f holomorphic_on (s - k)" and s: "open s" and k: "finite k"
-    shows "f holomorphic_on s"
-proof -
-  { fix z
-    assume "z \<in> s" and cdf: "\<And>x. x\<in>s - k \<Longrightarrow> f field_differentiable at x"
-    have "f field_differentiable at z"
-    proof (cases "z \<in> k")
-      case False then show ?thesis by (blast intro: cdf \<open>z \<in> s\<close>)
-    next
-      case True
-      with finite_set_avoid [OF k, of z]
-      obtain d where "d>0" and d: "\<And>x. \<lbrakk>x\<in>k; x \<noteq> z\<rbrakk> \<Longrightarrow> d \<le> dist z x"
-        by blast
-      obtain e where "e>0" and e: "ball z e \<subseteq> s"
-        using  s \<open>z \<in> s\<close> by (force simp add: open_contains_ball)
-      have fde: "continuous_on (ball z (min d e)) f"
-        by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI)
-      have "\<exists>g. \<forall>w \<in> ball z (min d e). (g has_field_derivative f w) (at w within ball z (min d e))"
-        apply (rule contour_integral_convex_primitive [OF convex_ball fde])
-        apply (rule Cauchy_theorem_triangle_cofinite [OF _ k])
-         apply (metis continuous_on_subset [OF fde] closed_segment_subset convex_ball starlike_convex_subset)
-        apply (rule cdf)
-        apply (metis Diff_iff Int_iff ball_min_Int bot_least contra_subsetD convex_ball e insert_subset
-               interior_mono interior_subset subset_hull)
-        done
-      then have "f holomorphic_on ball z (min d e)"
-        by (metis open_ball at_within_open derivative_is_holomorphic)
-      then show ?thesis
-        unfolding holomorphic_on_def
-        by (metis open_ball \<open>0 < d\<close> \<open>0 < e\<close> at_within_open centre_in_ball min_less_iff_conj)
-    qed
-  }
-  with holf s k show ?thesis
-    by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric])
-qed
-
-proposition Cauchy_integral_formula_convex:
-    assumes s: "convex s" and k: "finite k" and contf: "continuous_on s f"
-        and fcd: "(\<And>x. x \<in> interior s - k \<Longrightarrow> f field_differentiable at x)"
-        and z: "z \<in> interior s" and vpg: "valid_path \<gamma>"
-        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
-  apply (rule Cauchy_integral_formula_weak [OF s finite.emptyI contf])
-  apply (simp add: holomorphic_on_open [symmetric] field_differentiable_def)
-  using no_isolated_singularity [where s = "interior s"]
-  apply (metis k contf fcd holomorphic_on_open field_differentiable_def continuous_on_subset
-               has_field_derivative_at_within holomorphic_on_def interior_subset open_interior)
-  using assms
-  apply auto
-  done
-
-
-text\<open> Formula for higher derivatives.\<close>
-
-proposition Cauchy_has_contour_integral_higher_derivative_circlepath:
-  assumes contf: "continuous_on (cball z r) f"
-      and holf: "f holomorphic_on ball z r"
-      and w: "w \<in> ball z r"
-    shows "((\<lambda>u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \<i>) / (fact k) * (deriv ^^ k) f w))
-           (circlepath z r)"
-using w
-proof (induction k arbitrary: w)
-  case 0 then show ?case
-    using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm)
-next
-  case (Suc k)
-  have [simp]: "r > 0" using w
-    using ball_eq_empty by fastforce
-  have f: "continuous_on (path_image (circlepath z r)) f"
-    by (rule continuous_on_subset [OF contf]) (force simp add: cball_def sphere_def less_imp_le)
-  obtain X where X: "((\<lambda>u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)"
-    using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems]
-    by (auto simp: contour_integrable_on_def)
-  then have con: "contour_integral (circlepath z r) ((\<lambda>u. f u / (u - w) ^ Suc (Suc k))) = X"
-    by (rule contour_integral_unique)
-  have "\<And>n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)"
-    using Suc.prems assms has_field_derivative_higher_deriv by auto
-  then have dnf_diff: "\<And>n. (deriv ^^ n) f field_differentiable (at w)"
-    by (force simp add: field_differentiable_def)
-  have "deriv (\<lambda>w. complex_of_real (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w) w =
-          of_nat (Suc k) * contour_integral (circlepath z r) (\<lambda>u. f u / (u - w) ^ Suc (Suc k))"
-    by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems])
-  also have "... = of_nat (Suc k) * X"
-    by (simp only: con)
-  finally have "deriv (\<lambda>w. ((2 * pi) * \<i> / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" .
-  then have "((2 * pi) * \<i> / (fact k)) * deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X"
-    by (metis deriv_cmult dnf_diff)
-  then have "deriv (\<lambda>w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \<i> / (fact k))"
-    by (simp add: field_simps)
-  then show ?case
-  using of_nat_eq_0_iff X by fastforce
-qed
-
-proposition Cauchy_higher_derivative_integral_circlepath:
-  assumes contf: "continuous_on (cball z r) f"
-      and holf: "f holomorphic_on ball z r"
-      and w: "w \<in> ball z r"
-    shows "(\<lambda>u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)"
-           (is "?thes1")
-      and "(deriv ^^ k) f w = (fact k) / (2 * pi * \<i>) * contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k))"
-           (is "?thes2")
-proof -
-  have *: "((\<lambda>u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \<i> / (fact k) * (deriv ^^ k) f w)
-           (circlepath z r)"
-    using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms]
-    by simp
-  show ?thes1 using *
-    using contour_integrable_on_def by blast
-  show ?thes2
-    unfolding contour_integral_unique [OF *] by (simp add: divide_simps)
-qed
-
-corollary Cauchy_contour_integral_circlepath:
-  assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
-    shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^(Suc k)) = (2 * pi * \<i>) * (deriv ^^ k) f w / (fact k)"
-by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms])
-
-corollary Cauchy_contour_integral_circlepath_2:
-  assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \<in> ball z r"
-    shows "contour_integral(circlepath z r) (\<lambda>u. f u/(u - w)^2) = (2 * pi * \<i>) * deriv f w"
-  using Cauchy_contour_integral_circlepath [OF assms, of 1]
-  by (simp add: power2_eq_square)
-
-
-subsection\<open>A holomorphic function is analytic, i.e. has local power series.\<close>
-
-theorem holomorphic_power_series:
-  assumes holf: "f holomorphic_on ball z r"
-      and w: "w \<in> ball z r"
-    shows "((\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
-proof -
-  have fh': "f holomorphic_on cball z ((r + dist w z) / 2)"
-     apply (rule holomorphic_on_subset [OF holf])
-     apply (clarsimp simp del: divide_const_simps)
-     apply (metis add.commute dist_commute le_less_trans mem_ball real_gt_half_sum w)
-     done
-  \<comment>\<open>Replacing @{term r} and the original (weak) premises\<close>
-  obtain r where "0 < r" and holfc: "f holomorphic_on cball z r" and w: "w \<in> ball z r"
-    apply (rule that [of "(r + dist w z) / 2"])
-      apply (simp_all add: fh')
-     apply (metis add_0_iff ball_eq_empty dist_nz dist_self empty_iff not_less pos_add_strict w)
-    apply (metis add_less_cancel_right dist_commute mem_ball mult_2_right w)
-    done
-  then have holf: "f holomorphic_on ball z r" and contf: "continuous_on (cball z r) f"
-    using ball_subset_cball holomorphic_on_subset apply blast
-    by (simp add: holfc holomorphic_on_imp_continuous_on)
-  have cint: "\<And>k. (\<lambda>u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r"
-    apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
-    apply (simp add: \<open>0 < r\<close>)
-    done
-  obtain B where "0 < B" and B: "\<And>u. u \<in> cball z r \<Longrightarrow> norm(f u) \<le> B"
-    by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI)
-  obtain k where k: "0 < k" "k \<le> r" and wz_eq: "norm(w - z) = r - k"
-             and kle: "\<And>u. norm(u - z) = r \<Longrightarrow> k \<le> norm(u - w)"
-    apply (rule_tac k = "r - dist z w" in that)
-    using w
-    apply (auto simp: dist_norm norm_minus_commute)
-    by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
-  have *: "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>path_image (circlepath z r).
-                norm ((\<Sum>k<n. (w - z) ^ k * (f x / (x - z) ^ Suc k)) - f x / (x - w)) < e"
-          if "0 < e" for e
-  proof -
-    have rr: "0 \<le> (r - k) / r" "(r - k) / r < 1" using  k by auto
-    obtain n where n: "((r - k) / r) ^ n < e / B * k"
-      using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \<open>0 < e\<close> \<open>0 < B\<close> k by force
-    have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) - f u / (u - w)) < e"
-         if "n \<le> N" and r: "r = dist z u"  for N u
-    proof -
-      have N: "((r - k) / r) ^ N < e / B * k"
-        apply (rule le_less_trans [OF power_decreasing n])
-        using  \<open>n \<le> N\<close> k by auto
-      have u [simp]: "(u \<noteq> z) \<and> (u \<noteq> w)"
-        using \<open>0 < r\<close> r w by auto
-      have wzu_not1: "(w - z) / (u - z) \<noteq> 1"
-        by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w)
-      have "norm ((\<Sum>k<N. (w - z) ^ k * f u / (u - z) ^ Suc k) * (u - w) - f u)
-            = norm ((\<Sum>k<N. (((w - z) / (u - z)) ^ k)) * f u * (u - w) / (u - z) - f u)"
-        unfolding setsum_left_distrib setsum_divide_distrib power_divide by (simp add: algebra_simps)
-      also have "... = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)"
-        using \<open>0 < B\<close>
-        apply (auto simp: geometric_sum [OF wzu_not1])
-        apply (simp add: field_simps norm_mult [symmetric])
-        done
-      also have "... = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)"
-        using \<open>0 < r\<close> r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute)
-      also have "... = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)"
-        by (simp add: algebra_simps)
-      also have "... = norm (w - z) ^ N * norm (f u) / r ^ N"
-        by (simp add: norm_mult norm_power norm_minus_commute)
-      also have "... \<le> (((r - k)/r)^N) * B"
-        using \<open>0 < r\<close> w k
-        apply (simp add: divide_simps)
-        apply (rule mult_mono [OF power_mono])
-        apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r)
-        done
-      also have "... < e * k"
-        using \<open>0 < B\<close> N by (simp add: divide_simps)
-      also have "... \<le> e * norm (u - w)"
-        using r kle \<open>0 < e\<close> by (simp add: dist_commute dist_norm)
-      finally show ?thesis
-        by (simp add: divide_simps norm_divide del: power_Suc)
-    qed
-    with \<open>0 < r\<close> show ?thesis
-      by (auto simp: mult_ac less_imp_le eventually_sequentially Ball_def)
-  qed
-  have eq: "\<forall>\<^sub>F x in sequentially.
-             contour_integral (circlepath z r) (\<lambda>u. \<Sum>k<x. (w - z) ^ k * (f u / (u - z) ^ Suc k)) =
-             (\<Sum>k<x. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z) ^ k)"
-    apply (rule eventuallyI)
-    apply (subst contour_integral_setsum, simp)
-    using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps)
-    apply (simp only: contour_integral_lmul cint algebra_simps)
-    done
-  have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
-        sums contour_integral (circlepath z r) (\<lambda>u. f u/(u - w))"
-    unfolding sums_def
-    apply (rule Lim_transform_eventually [OF eq])
-    apply (rule contour_integral_uniform_limit_circlepath [OF eventuallyI *])
-    apply (rule contour_integrable_setsum, simp)
-    apply (rule contour_integrable_lmul)
-    apply (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf])
-    using \<open>0 < r\<close>
-    apply auto
-    done
-  then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u/(u - z)^(Suc k)) * (w - z)^k)
-             sums (2 * of_real pi * \<i> * f w)"
-    using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]])
-  then have "(\<lambda>k. contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc k) * (w - z)^k / (\<i> * (of_real pi * 2)))
-            sums ((2 * of_real pi * \<i> * f w) / (\<i> * (complex_of_real pi * 2)))"
-    by (rule sums_divide)
-  then have "(\<lambda>n. (w - z) ^ n * contour_integral (circlepath z r) (\<lambda>u. f u / (u - z) ^ Suc n) / (\<i> * (of_real pi * 2)))
-            sums f w"
-    by (simp add: field_simps)
-  then show ?thesis
-    by (simp add: field_simps \<open>0 < r\<close> Cauchy_higher_derivative_integral_circlepath [OF contf holf])
-qed
-
-
-subsection\<open>The Liouville theorem and the Fundamental Theorem of Algebra.\<close>
-
-text\<open> These weak Liouville versions don't even need the derivative formula.\<close>
-
-lemma Liouville_weak_0:
-  assumes holf: "f holomorphic_on UNIV" and inf: "(f \<longlongrightarrow> 0) at_infinity"
-    shows "f z = 0"
-proof (rule ccontr)
-  assume fz: "f z \<noteq> 0"
-  with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"]
-  obtain B where B: "\<And>x. B \<le> cmod x \<Longrightarrow> norm (f x) * 2 < cmod (f z)"
-    by (auto simp: dist_norm)
-  define R where "R = 1 + \<bar>B\<bar> + norm z"
-  have "R > 0" unfolding R_def
-  proof -
-    have "0 \<le> cmod z + \<bar>B\<bar>"
-      by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def)
-    then show "0 < 1 + \<bar>B\<bar> + cmod z"
-      by linarith
-  qed
-  have *: "((\<lambda>u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \<i> * f z) (circlepath z R)"
-    apply (rule Cauchy_integral_circlepath)
-    using \<open>R > 0\<close> apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+
-    done
-  have "cmod (x - z) = R \<Longrightarrow> cmod (f x) * 2 \<le> cmod (f z)" for x
-    apply (simp add: R_def)
-    apply (rule less_imp_le)
-    apply (rule B)
-    using norm_triangle_ineq4 [of x z]
-    apply (auto simp:)
-    done
-  with  \<open>R > 0\<close> fz show False
-    using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"]
-    by (auto simp: norm_mult norm_divide divide_simps)
-qed
-
-proposition Liouville_weak:
-  assumes "f holomorphic_on UNIV" and "(f \<longlongrightarrow> l) at_infinity"
-    shows "f z = l"
-  using Liouville_weak_0 [of "\<lambda>z. f z - l"]
-  by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero)
-
-
-proposition Liouville_weak_inverse:
-  assumes "f holomorphic_on UNIV" and unbounded: "\<And>B. eventually (\<lambda>x. norm (f x) \<ge> B) at_infinity"
-    obtains z where "f z = 0"
-proof -
-  { assume f: "\<And>z. f z \<noteq> 0"
-    have 1: "(\<lambda>x. 1 / f x) holomorphic_on UNIV"
-      by (simp add: holomorphic_on_divide holomorphic_on_const assms f)
-    have 2: "((\<lambda>x. 1 / f x) \<longlongrightarrow> 0) at_infinity"
-      apply (rule tendstoI [OF eventually_mono])
-      apply (rule_tac B="2/e" in unbounded)
-      apply (simp add: dist_norm norm_divide divide_simps mult_ac)
-      done
-    have False
-      using Liouville_weak_0 [OF 1 2] f by simp
-  }
-  then show ?thesis
-    using that by blast
-qed
-
-
-text\<open> In particular we get the Fundamental Theorem of Algebra.\<close>
-
-theorem fundamental_theorem_of_algebra:
-    fixes a :: "nat \<Rightarrow> complex"
-  assumes "a 0 = 0 \<or> (\<exists>i \<in> {1..n}. a i \<noteq> 0)"
-  obtains z where "(\<Sum>i\<le>n. a i * z^i) = 0"
-using assms
-proof (elim disjE bexE)
-  assume "a 0 = 0" then show ?thesis
-    by (auto simp: that [of 0])
-next
-  fix i
-  assume i: "i \<in> {1..n}" and nz: "a i \<noteq> 0"
-  have 1: "(\<lambda>z. \<Sum>i\<le>n. a i * z^i) holomorphic_on UNIV"
-    by (rule holomorphic_intros)+
-  show ?thesis
-    apply (rule Liouville_weak_inverse [OF 1])
-    apply (rule polyfun_extremal)
-    apply (rule nz)
-    using i that
-    apply (auto simp:)
-    done
-qed
-
-
-subsection\<open> Weierstrass convergence theorem.\<close>
-
-proposition holomorphic_uniform_limit:
-  assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and> (f n) holomorphic_on ball z r) F"
-      and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
-      and F:  "~ trivial_limit F"
-  obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r"
-proof (cases r "0::real" rule: linorder_cases)
-  case less then show ?thesis by (force simp add: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that)
-next
-  case equal then show ?thesis
-    by (force simp add: holomorphic_on_def continuous_on_sing intro: that)
-next
-  case greater
-  have contg: "continuous_on (cball z r) g"
-    using cont
-    by (fastforce simp: eventually_conj_iff dist_norm intro: eventually_mono [OF lim] continuous_uniform_limit [OF F])
-  have 1: "continuous_on (path_image (circlepath z r)) (\<lambda>x. 1 / (2 * complex_of_real pi * \<i>) * g x)"
-    apply (rule continuous_intros continuous_on_subset [OF contg])+
-    using \<open>0 < r\<close> by auto
-  have 2: "((\<lambda>u. 1 / (2 * of_real pi * \<i>) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)"
-       if w: "w \<in> ball z r" for w
-  proof -
-    define d where "d = (r - norm(w - z))"
-    have "0 < d"  "d \<le> r" using w by (auto simp: norm_minus_commute d_def dist_norm)
-    have dle: "\<And>u. cmod (z - u) = r \<Longrightarrow> d \<le> cmod (u - w)"
-      unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
-    have ev_int: "\<forall>\<^sub>F n in F. (\<lambda>u. f n u / (u - w)) contour_integrable_on circlepath z r"
-      apply (rule eventually_mono [OF cont])
-      using w
-      apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified])
-      done
-    have ev_less: "\<forall>\<^sub>F n in F. \<forall>x\<in>path_image (circlepath z r). cmod (f n x / (x - w) - g x / (x - w)) < e"
-         if "e > 0" for e
-      using greater \<open>0 < d\<close> \<open>0 < e\<close>
-      apply (simp add: norm_divide diff_divide_distrib [symmetric] divide_simps)
-      apply (rule_tac e1="e * d" in eventually_mono [OF lim])
-      apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+
-      done
-    have g_cint: "(\<lambda>u. g u/(u - w)) contour_integrable_on circlepath z r"
-      by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
-    have cif_tends_cig: "((\<lambda>n. contour_integral(circlepath z r) (\<lambda>u. f n u / (u - w))) \<longlongrightarrow> contour_integral(circlepath z r) (\<lambda>u. g u/(u - w))) F"
-      by (rule contour_integral_uniform_limit_circlepath [OF ev_int ev_less F \<open>0 < r\<close>])
-    have f_tends_cig: "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> contour_integral (circlepath z r) (\<lambda>u. g u / (u - w))) F"
-      apply (rule Lim_transform_eventually [where f = "\<lambda>n. contour_integral (circlepath z r) (\<lambda>u. f n u/(u - w))"])
-      apply (rule eventually_mono [OF cont])
-      apply (rule contour_integral_unique [OF Cauchy_integral_circlepath])
-      using w
-      apply (auto simp: norm_minus_commute dist_norm cif_tends_cig)
-      done
-    have "((\<lambda>n. 2 * of_real pi * \<i> * f n w) \<longlongrightarrow> 2 * of_real pi * \<i> * g w) F"
-      apply (rule tendsto_mult_left [OF tendstoI])
-      apply (rule eventually_mono [OF lim], assumption)
-      using w
-      apply (force simp add: dist_norm)
-      done
-    then have "((\<lambda>u. g u / (u - w)) has_contour_integral 2 * of_real pi * \<i> * g w) (circlepath z r)"
-      using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w
-      by (force simp add: dist_norm)
-    then have "((\<lambda>u. g u / (2 * of_real pi * \<i> * (u - w))) has_contour_integral g w) (circlepath z r)"
-      using has_contour_integral_div [where c = "2 * of_real pi * \<i>"]
-      by (force simp add: field_simps)
-    then show ?thesis
-      by (simp add: dist_norm)
-  qed
-  show ?thesis
-    using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified]
-    by (fastforce simp add: holomorphic_on_open contg intro: that)
-qed
-
-
-text\<open> Version showing that the limit is the limit of the derivatives.\<close>
-
-proposition has_complex_derivative_uniform_limit:
-  fixes z::complex
-  assumes cont: "eventually (\<lambda>n. continuous_on (cball z r) (f n) \<and>
-                               (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F"
-      and lim: "\<And>e. 0 < e \<Longrightarrow> eventually (\<lambda>n. \<forall>x \<in> cball z r. norm(f n x - g x) < e) F"
-      and F:  "~ trivial_limit F" and "0 < r"
-  obtains g' where
-      "continuous_on (cball z r) g"
-      "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
-proof -
-  let ?conint = "contour_integral (circlepath z r)"
-  have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r"
-    by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] lim F];
-             auto simp: holomorphic_on_open field_differentiable_def)+
-  then obtain g' where g': "\<And>x. x \<in> ball z r \<Longrightarrow> (g has_field_derivative g' x) (at x)"
-    using DERIV_deriv_iff_has_field_derivative
-    by (fastforce simp add: holomorphic_on_open)
-  then have derg: "\<And>x. x \<in> ball z r \<Longrightarrow> deriv g x = g' x"
-    by (simp add: DERIV_imp_deriv)
-  have tends_f'n_g': "((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F" if w: "w \<in> ball z r" for w
-  proof -
-    have eq_f': "?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \<i>)"
-             if cont_fn: "continuous_on (cball z r) (f n)"
-             and fnd: "\<And>w. w \<in> ball z r \<Longrightarrow> (f n has_field_derivative f' n w) (at w)" for n
-    proof -
-      have hol_fn: "f n holomorphic_on ball z r"
-        using fnd by (force simp add: holomorphic_on_open)
-      have "(f n has_field_derivative 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)) (at w)"
-        by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w])
-      then have f': "f' n w = 1 / (2 * of_real pi * \<i>) * ?conint (\<lambda>u. f n u / (u - w)\<^sup>2)"
-        using DERIV_unique [OF fnd] w by blast
-      show ?thesis
-        by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] divide_simps)
-    qed
-    define d where "d = (r - norm(w - z))^2"
-    have "d > 0"
-      using w by (simp add: dist_commute dist_norm d_def)
-    have dle: "\<And>y. r = cmod (z - y) \<Longrightarrow> d \<le> cmod ((y - w)\<^sup>2)"
-      apply (simp add: d_def norm_power)
-      apply (rule power_mono)
-      apply (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute)
-      apply (metis diff_ge_0_iff_ge dist_commute dist_norm less_eq_real_def mem_ball w)
-      done
-    have 1: "\<forall>\<^sub>F n in F. (\<lambda>x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r"
-      by (force simp add: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont])
-    have 2: "0 < e \<Longrightarrow> \<forall>\<^sub>F n in F. \<forall>x \<in> path_image (circlepath z r). cmod (f n x / (x - w)\<^sup>2 - g x / (x - w)\<^sup>2) < e" for e
-      using \<open>r > 0\<close>
-      apply (simp add: diff_divide_distrib [symmetric] norm_divide divide_simps sphere_def)
-      apply (rule eventually_mono [OF lim, of "e*d"])
-      apply (simp add: \<open>0 < d\<close>)
-      apply (force simp add: dist_norm dle intro: less_le_trans)
-      done
-    have "((\<lambda>n. contour_integral (circlepath z r) (\<lambda>x. f n x / (x - w)\<^sup>2))
-             \<longlongrightarrow> contour_integral (circlepath z r) ((\<lambda>x. g x / (x - w)\<^sup>2))) F"
-      by (rule Cauchy_Integral_Thm.contour_integral_uniform_limit_circlepath [OF 1 2 F \<open>0 < r\<close>])
-    then have tendsto_0: "((\<lambda>n. 1 / (2 * of_real pi * \<i>) * (?conint (\<lambda>x. f n x / (x - w)\<^sup>2) - ?conint (\<lambda>x. g x / (x - w)\<^sup>2))) \<longlongrightarrow> 0) F"
-      using Lim_null by (force intro!: tendsto_mult_right_zero)
-    have "((\<lambda>n. f' n w - g' w) \<longlongrightarrow> 0) F"
-      apply (rule Lim_transform_eventually [OF _ tendsto_0])
-      apply (force simp add: divide_simps intro: eq_f' eventually_mono [OF cont])
-      done
-    then show ?thesis using Lim_null by blast
-  qed
-  obtain g' where "\<And>w. w \<in> ball z r \<Longrightarrow> (g has_field_derivative (g' w)) (at w) \<and> ((\<lambda>n. f' n w) \<longlongrightarrow> g' w) F"
-      by (blast intro: tends_f'n_g' g' )
-  then show ?thesis using g
-    using that by blast
-qed
-
-
-subsection\<open>Some more simple/convenient versions for applications.\<close>
-
-lemma holomorphic_uniform_sequence:
-  assumes s: "open s"
-      and hol_fn: "\<And>n. (f n) holomorphic_on s"
-      and to_g: "\<And>x. x \<in> s
-                     \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
-                             (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
-  shows "g holomorphic_on s"
-proof -
-  have "\<exists>f'. (g has_field_derivative f') (at z)" if "z \<in> s" for z
-  proof -
-    obtain r where "0 < r" and r: "cball z r \<subseteq> s"
-               and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
-      using to_g [OF \<open>z \<in> s\<close>] by blast
-    have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and> f n holomorphic_on ball z r"
-      apply (intro eventuallyI conjI)
-      using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r apply blast
-      apply (metis hol_fn holomorphic_on_subset interior_cball interior_subset r)
-      done
-    show ?thesis
-      apply (rule holomorphic_uniform_limit [OF *])
-      using \<open>0 < r\<close> centre_in_ball fg
-      apply (auto simp: holomorphic_on_open)
-      done
-  qed
-  with s show ?thesis
-    by (simp add: holomorphic_on_open)
-qed
-
-lemma has_complex_derivative_uniform_sequence:
-  fixes s :: "complex set"
-  assumes s: "open s"
-      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> ((f n) has_field_derivative f' n x) (at x)"
-      and to_g: "\<And>x. x \<in> s
-             \<Longrightarrow> \<exists>d. 0 < d \<and> cball x d \<subseteq> s \<and>
-                     (\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball x d. norm(f n y - g y) < e) sequentially)"
-  shows "\<exists>g'. \<forall>x \<in> s. (g has_field_derivative g' x) (at x) \<and> ((\<lambda>n. f' n x) \<longlongrightarrow> g' x) sequentially"
-proof -
-  have y: "\<exists>y. (g has_field_derivative y) (at z) \<and> (\<lambda>n. f' n z) \<longlonglongrightarrow> y" if "z \<in> s" for z
-  proof -
-    obtain r where "0 < r" and r: "cball z r \<subseteq> s"
-               and fg: "\<forall>e. 0 < e \<longrightarrow> eventually (\<lambda>n. \<forall>y \<in> cball z r. norm(f n y - g y) < e) sequentially"
-      using to_g [OF \<open>z \<in> s\<close>] by blast
-    have *: "\<forall>\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \<and>
-                                   (\<forall>w \<in> ball z r. ((f n) has_field_derivative (f' n w)) (at w))"
-      apply (intro eventuallyI conjI)
-      apply (meson hfd holomorphic_on_imp_continuous_on holomorphic_on_open holomorphic_on_subset r s)
-      using ball_subset_cball hfd r apply blast
-      done
-    show ?thesis
-      apply (rule has_complex_derivative_uniform_limit [OF *, of g])
-      using \<open>0 < r\<close> centre_in_ball fg
-      apply force+
-      done
-  qed
-  show ?thesis
-    by (rule bchoice) (blast intro: y)
-qed
-
-
-subsection\<open>On analytic functions defined by a series.\<close>
-
-lemma series_and_derivative_comparison:
-  fixes s :: "complex set"
-  assumes s: "open s"
-      and h: "summable h"
-      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-      and to_g: "\<And>n x. \<lbrakk>N \<le> n; x \<in> s\<rbrakk> \<Longrightarrow> norm(f n x) \<le> h n"
-  obtains g g' where "\<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
-proof -
-  obtain g where g: "\<And>e. e>0 \<Longrightarrow> \<exists>N. \<forall>n x. N \<le> n \<and> x \<in> s \<longrightarrow> dist (\<Sum>n<n. f n x) (g x) < e"
-    using series_comparison_uniform [OF h to_g, of N s] by force
-  have *: "\<exists>d>0. cball x d \<subseteq> s \<and> (\<forall>e>0. \<forall>\<^sub>F n in sequentially. \<forall>y\<in>cball x d. cmod ((\<Sum>a<n. f a y) - g y) < e)"
-         if "x \<in> s" for x
-  proof -
-    obtain d where "d>0" and d: "cball x d \<subseteq> s"
-      using open_contains_cball [of "s"] \<open>x \<in> s\<close> s by blast
-    then show ?thesis
-      apply (rule_tac x=d in exI)
-      apply (auto simp: dist_norm eventually_sequentially)
-      apply (metis g contra_subsetD dist_norm)
-      done
-  qed
-  have "(\<forall>x\<in>s. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> g x)"
-    using g by (force simp add: lim_sequentially)
-  moreover have "\<exists>g'. \<forall>x\<in>s. (g has_field_derivative g' x) (at x) \<and> (\<lambda>n. \<Sum>i<n. f' i x) \<longlonglongrightarrow> g' x"
-    by (rule has_complex_derivative_uniform_sequence [OF s]) (auto intro: * hfd DERIV_setsum)+
-  ultimately show ?thesis
-    by (force simp add: sums_def  conj_commute intro: that)
-qed
-
-text\<open>A version where we only have local uniform/comparative convergence.\<close>
-
-lemma series_and_derivative_comparison_local:
-  fixes s :: "complex set"
-  assumes s: "open s"
-      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-      and to_g: "\<And>x. x \<in> s \<Longrightarrow>
-                      \<exists>d h N. 0 < d \<and> summable h \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> norm(f n y) \<le> h n)"
-  shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
-proof -
-  have "\<exists>y. (\<lambda>n. f n z) sums (\<Sum>n. f n z) \<and> (\<lambda>n. f' n z) sums y \<and> ((\<lambda>x. \<Sum>n. f n x) has_field_derivative y) (at z)"
-       if "z \<in> s" for z
-  proof -
-    obtain d h N where "0 < d" "summable h" and le_h: "\<And>n y. \<lbrakk>N \<le> n; y \<in> ball z d\<rbrakk> \<Longrightarrow> norm(f n y) \<le> h n"
-      using to_g \<open>z \<in> s\<close> by meson
-    then obtain r where "r>0" and r: "ball z r \<subseteq> ball z d \<inter> s" using \<open>z \<in> s\<close> s
-      by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq)
-    have 1: "open (ball z d \<inter> s)"
-      by (simp add: open_Int s)
-    have 2: "\<And>n x. x \<in> ball z d \<inter> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-      by (auto simp: hfd)
-    obtain g g' where gg': "\<forall>x \<in> ball z d \<inter> s. ((\<lambda>n. f n x) sums g x) \<and>
-                                    ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
-      by (auto intro: le_h series_and_derivative_comparison [OF 1 \<open>summable h\<close> hfd])
-    then have "(\<lambda>n. f' n z) sums g' z"
-      by (meson \<open>0 < r\<close> centre_in_ball contra_subsetD r)
-    moreover have "(\<lambda>n. f n z) sums (\<Sum>n. f n z)"
-      by (metis summable_comparison_test' summable_sums centre_in_ball \<open>0 < d\<close> \<open>summable h\<close> le_h)
-    moreover have "((\<lambda>x. \<Sum>n. f n x) has_field_derivative g' z) (at z)"
-      apply (rule_tac f=g in DERIV_transform_at [OF _ \<open>0 < r\<close>])
-      apply (simp add: gg' \<open>z \<in> s\<close> \<open>0 < d\<close>)
-      apply (metis (full_types) contra_subsetD dist_commute gg' mem_ball r sums_unique)
-      done
-    ultimately show ?thesis by auto
-  qed
-  then show ?thesis
-    by (rule_tac x="\<lambda>x. suminf  (\<lambda>n. f n x)" in exI) meson
-qed
-
-
-text\<open>Sometimes convenient to compare with a complex series of positive reals. (?)\<close>
-
-lemma series_and_derivative_comparison_complex:
-  fixes s :: "complex set"
-  assumes s: "open s"
-      and hfd: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)"
-      and to_g: "\<And>x. x \<in> s \<Longrightarrow>
-                      \<exists>d h N. 0 < d \<and> summable h \<and> range h \<subseteq> nonneg_Reals \<and> (\<forall>n y. N \<le> n \<and> y \<in> ball x d \<longrightarrow> cmod(f n y) \<le> cmod (h n))"
-  shows "\<exists>g g'. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((\<lambda>n. f' n x) sums g' x) \<and> (g has_field_derivative g' x) (at x)"
-apply (rule series_and_derivative_comparison_local [OF s hfd], assumption)
-apply (frule to_g)
-apply (erule ex_forward)
-apply (erule exE)
-apply (rule_tac x="Re o h" in exI)
-apply (erule ex_forward)
-apply (simp add: summable_Re o_def )
-apply (elim conjE all_forward)
-apply (simp add: nonneg_Reals_cmod_eq_Re image_subset_iff)
-done
-
-
-text\<open>In particular, a power series is analytic inside circle of convergence.\<close>
-
-lemma power_series_and_derivative_0:
-  fixes a :: "nat \<Rightarrow> complex" and r::real
-  assumes "summable (\<lambda>n. a n * r^n)"
-    shows "\<exists>g g'. \<forall>z. cmod z < r \<longrightarrow>
-             ((\<lambda>n. a n * z^n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * z^(n - 1)) sums g' z) \<and> (g has_field_derivative g' z) (at z)"
-proof (cases "0 < r")
-  case True
-    have der: "\<And>n z. ((\<lambda>x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)"
-      by (rule derivative_eq_intros | simp)+
-    have y_le: "\<lbrakk>cmod (z - y) * 2 < r - cmod z\<rbrakk> \<Longrightarrow> cmod y \<le> cmod (of_real r + of_real (cmod z)) / 2" for z y
-      using \<open>r > 0\<close>
-      apply (auto simp: algebra_simps norm_mult norm_divide norm_power of_real_add [symmetric] simp del: of_real_add)
-      using norm_triangle_ineq2 [of y z]
-      apply (simp only: diff_le_eq norm_minus_commute mult_2)
-      done
-    have "summable (\<lambda>n. a n * complex_of_real r ^ n)"
-      using assms \<open>r > 0\<close> by simp
-    moreover have "\<And>z. cmod z < r \<Longrightarrow> cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)"
-      using \<open>r > 0\<close>
-      by (simp add: of_real_add [symmetric] del: of_real_add)
-    ultimately have sum: "\<And>z. cmod z < r \<Longrightarrow> summable (\<lambda>n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)"
-      by (rule power_series_conv_imp_absconv_weak)
-    have "\<exists>g g'. \<forall>z \<in> ball 0 r. (\<lambda>n.  (a n) * z ^ n) sums g z \<and>
-               (\<lambda>n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \<and> (g has_field_derivative g' z) (at z)"
-      apply (rule series_and_derivative_comparison_complex [OF open_ball der])
-      apply (rule_tac x="(r - norm z)/2" in exI)
-      apply (simp add: dist_norm)
-      apply (rule_tac x="\<lambda>n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI)
-      using \<open>r > 0\<close>
-      apply (auto simp: sum nonneg_Reals_divide_I)
-      apply (rule_tac x=0 in exI)
-      apply (force simp: norm_mult norm_divide norm_power intro!: mult_left_mono power_mono y_le)
-      done
-  then show ?thesis
-    by (simp add: dist_0_norm ball_def)
-next
-  case False then show ?thesis
-    apply (simp add: not_less)
-    using less_le_trans norm_not_less_zero by blast
-qed
-
-proposition power_series_and_derivative:
-  fixes a :: "nat \<Rightarrow> complex" and r::real
-  assumes "summable (\<lambda>n. a n * r^n)"
-    obtains g g' where "\<forall>z \<in> ball w r.
-             ((\<lambda>n. a n * (z - w) ^ n) sums g z) \<and> ((\<lambda>n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \<and>
-              (g has_field_derivative g' z) (at z)"
-  using power_series_and_derivative_0 [OF assms]
-  apply clarify
-  apply (rule_tac g="(\<lambda>z. g(z - w))" in that)
-  using DERIV_shift [where z="-w"]
-  apply (auto simp: norm_minus_commute Ball_def dist_norm)
-  done
-
-proposition power_series_holomorphic:
-  assumes "\<And>w. w \<in> ball z r \<Longrightarrow> ((\<lambda>n. a n*(w - z)^n) sums f w)"
-    shows "f holomorphic_on ball z r"
-proof -
-  have "\<exists>f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w
-  proof -
-    have inb: "z + complex_of_real ((dist z w + r) / 2) \<in> ball z r"
-    proof -
-      have wz: "cmod (w - z) < r" using w
-        by (auto simp: divide_simps dist_norm norm_minus_commute)
-      then have "0 \<le> r"
-        by (meson less_eq_real_def norm_ge_zero order_trans)
-      show ?thesis
-        using w by (simp add: dist_norm \<open>0\<le>r\<close> of_real_add [symmetric] del: of_real_add)
-    qed
-    have sum: "summable (\<lambda>n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))"
-      using assms [OF inb] by (force simp add: summable_def dist_norm)
-    obtain g g' where gg': "\<And>u. u \<in> ball z ((cmod (z - w) + r) / 2) \<Longrightarrow>
-                               (\<lambda>n. a n * (u - z) ^ n) sums g u \<and>
-                               (\<lambda>n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \<and> (g has_field_derivative g' u) (at u)"
-      by (rule power_series_and_derivative [OF sum, of z]) fastforce
-    have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u
-    proof -
-      have less: "cmod (z - u) * 2 < cmod (z - w) + r"
-        using that dist_triangle2 [of z u w]
-        by (simp add: dist_norm [symmetric] algebra_simps)
-      show ?thesis
-        apply (rule sums_unique2 [of "\<lambda>n. a n*(u - z)^n"])
-        using gg' [of u] less w
-        apply (auto simp: assms dist_norm)
-        done
-    qed
-    show ?thesis
-      apply (rule_tac x="g' w" in exI)
-      apply (rule DERIV_transform_at [where f=g and d="(r - norm(z - w))/2"])
-      using w gg' [of w]
-      apply (auto simp: dist_norm)
-      done
-  qed
-  then show ?thesis by (simp add: holomorphic_on_open)
-qed
-
-corollary holomorphic_iff_power_series:
-     "f holomorphic_on ball z r \<longleftrightarrow>
-      (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
-  apply (intro iffI ballI)
-   using holomorphic_power_series  apply force
-  apply (rule power_series_holomorphic [where a = "\<lambda>n. (deriv ^^ n) f z / (fact n)"])
-  apply force
-  done
-
-corollary power_series_analytic:
-     "(\<And>w. w \<in> ball z r \<Longrightarrow> (\<lambda>n. a n*(w - z)^n) sums f w) \<Longrightarrow> f analytic_on ball z r"
-  by (force simp add: analytic_on_open intro!: power_series_holomorphic)
-
-corollary analytic_iff_power_series:
-     "f analytic_on ball z r \<longleftrightarrow>
-      (\<forall>w \<in> ball z r. (\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)"
-  by (simp add: analytic_on_open holomorphic_iff_power_series)
-
-
-subsection\<open>Equality between holomorphic functions, on open ball then connected set.\<close>
-
-lemma holomorphic_fun_eq_on_ball:
-   "\<lbrakk>f holomorphic_on ball z r; g holomorphic_on ball z r;
-     w \<in> ball z r;
-     \<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z\<rbrakk>
-     \<Longrightarrow> f w = g w"
-  apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
-  apply (auto simp: holomorphic_iff_power_series)
-  done
-
-lemma holomorphic_fun_eq_0_on_ball:
-   "\<lbrakk>f holomorphic_on ball z r;  w \<in> ball z r;
-     \<And>n. (deriv ^^ n) f z = 0\<rbrakk>
-     \<Longrightarrow> f w = 0"
-  apply (rule sums_unique2 [of "\<lambda>n. (deriv ^^ n) f z / (fact n) * (w - z)^n"])
-  apply (auto simp: holomorphic_iff_power_series)
-  done
-
-lemma holomorphic_fun_eq_0_on_connected:
-  assumes holf: "f holomorphic_on s" and "open s"
-      and cons: "connected s"
-      and der: "\<And>n. (deriv ^^ n) f z = 0"
-      and "z \<in> s" "w \<in> s"
-    shows "f w = 0"
-proof -
-  have *: "\<And>x e. \<lbrakk> \<forall>xa. (deriv ^^ xa) f x = 0;  ball x e \<subseteq> s\<rbrakk>
-           \<Longrightarrow> ball x e \<subseteq> (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
-    apply auto
-    apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv])
-    apply (rule holomorphic_on_subset [OF holf], simp_all)
-    by (metis funpow_add o_apply)
-  have 1: "openin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
-    apply (rule open_subset, force)
-    using \<open>open s\<close>
-    apply (simp add: open_contains_ball Ball_def)
-    apply (erule all_forward)
-    using "*" by auto blast+
-  have 2: "closedin (subtopology euclidean s) (\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0})"
-    using assms
-    by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv)
-  obtain e where "e>0" and e: "ball w e \<subseteq> s" using openE [OF \<open>open s\<close> \<open>w \<in> s\<close>] .
-  then have holfb: "f holomorphic_on ball w e"
-    using holf holomorphic_on_subset by blast
-  have 3: "(\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}) = s \<Longrightarrow> f w = 0"
-    using \<open>e>0\<close> e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb])
-  show ?thesis
-    using cons der \<open>z \<in> s\<close>
-    apply (simp add: connected_clopen)
-    apply (drule_tac x="\<Inter>n. {w \<in> s. (deriv ^^ n) f w = 0}" in spec)
-    apply (auto simp: 1 2 3)
-    done
-qed
-
-lemma holomorphic_fun_eq_on_connected:
-  assumes "f holomorphic_on s" "g holomorphic_on s" and "open s"  "connected s"
-      and "\<And>n. (deriv ^^ n) f z = (deriv ^^ n) g z"
-      and "z \<in> s" "w \<in> s"
-    shows "f w = g w"
-  apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>x. f x - g x" s z, simplified])
-  apply (intro assms holomorphic_intros)
-  using assms apply simp_all
-  apply (subst higher_deriv_diff, auto)
-  done
-
-lemma holomorphic_fun_eq_const_on_connected:
-  assumes holf: "f holomorphic_on s" and "open s"
-      and cons: "connected s"
-      and der: "\<And>n. 0 < n \<Longrightarrow> (deriv ^^ n) f z = 0"
-      and "z \<in> s" "w \<in> s"
-    shows "f w = f z"
-  apply (rule holomorphic_fun_eq_0_on_connected [of "\<lambda>w. f w - f z" s z, simplified])
-  apply (intro assms holomorphic_intros)
-  using assms apply simp_all
-  apply (subst higher_deriv_diff)
-  apply (intro holomorphic_intros | simp)+
-  done
-
-
-subsection\<open>Some basic lemmas about poles/singularities.\<close>
-
-lemma pole_lemma:
-  assumes holf: "f holomorphic_on s" and a: "a \<in> interior s"
-    shows "(\<lambda>z. if z = a then deriv f a
-                 else (f z - f a) / (z - a)) holomorphic_on s" (is "?F holomorphic_on s")
-proof -
-  have F1: "?F field_differentiable (at u within s)" if "u \<in> s" "u \<noteq> a" for u
-  proof -
-    have fcd: "f field_differentiable at u within s"
-      using holf holomorphic_on_def by (simp add: \<open>u \<in> s\<close>)
-    have cd: "(\<lambda>z. (f z - f a) / (z - a)) field_differentiable at u within s"
-      by (rule fcd derivative_intros | simp add: that)+
-    have "0 < dist a u" using that dist_nz by blast
-    then show ?thesis
-      by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \<open>u \<in> s\<close>)
-  qed
-  have F2: "?F field_differentiable at a" if "0 < e" "ball a e \<subseteq> s" for e
-  proof -
-    have holfb: "f holomorphic_on ball a e"
-      by (rule holomorphic_on_subset [OF holf \<open>ball a e \<subseteq> s\<close>])
-    have 2: "?F holomorphic_on ball a e - {a}"
-      apply (rule holomorphic_on_subset [where s = "s - {a}"])
-      apply (simp add: holomorphic_on_def field_differentiable_def [symmetric])
-      using mem_ball that
-      apply (auto intro: F1 field_differentiable_within_subset)
-      done
-    have "isCont (\<lambda>z. if z = a then deriv f a else (f z - f a) / (z - a)) x"
-            if "dist a x < e" for x
-    proof (cases "x=a")
-      case True then show ?thesis
-      using holfb \<open>0 < e\<close>
-      apply (simp add: holomorphic_on_open field_differentiable_def [symmetric])
-      apply (drule_tac x=a in bspec)
-      apply (auto simp: DERIV_deriv_iff_field_differentiable [symmetric] continuous_at DERIV_iff2
-                elim: rev_iffD1 [OF _ LIM_equal])
-      done
-    next
-      case False with 2 that show ?thesis
-        by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at)
-    qed
-    then have 1: "continuous_on (ball a e) ?F"
-      by (clarsimp simp:  continuous_on_eq_continuous_at)
-    have "?F holomorphic_on ball a e"
-      by (auto intro: no_isolated_singularity [OF 1 2])
-    with that show ?thesis
-      by (simp add: holomorphic_on_open field_differentiable_def [symmetric]
-                    field_differentiable_at_within)
-  qed
-  show ?thesis
-  proof
-    fix x assume "x \<in> s" show "?F field_differentiable at x within s"
-    proof (cases "x=a")
-      case True then show ?thesis
-      using a by (auto simp: mem_interior intro: field_differentiable_at_within F2)
-    next
-      case False with F1 \<open>x \<in> s\<close>
-      show ?thesis by blast
-    qed
-  qed
-qed
-
-proposition pole_theorem:
-  assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
-      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
-    shows "(\<lambda>z. if z = a then deriv g a
-                 else f z - g a/(z - a)) holomorphic_on s"
-  using pole_lemma [OF holg a]
-  by (rule holomorphic_transform) (simp add: eq divide_simps)
-
-lemma pole_lemma_open:
-  assumes "f holomorphic_on s" "open s"
-    shows "(\<lambda>z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on s"
-proof (cases "a \<in> s")
-  case True with assms interior_eq pole_lemma
-    show ?thesis by fastforce
-next
-  case False with assms show ?thesis
-    apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify)
-    apply (rule field_differentiable_transform_within [where f = "\<lambda>z. (f z - f a)/(z - a)" and d = 1])
-    apply (rule derivative_intros | force)+
-    done
-qed
-
-proposition pole_theorem_open:
-  assumes holg: "g holomorphic_on s" and s: "open s"
-      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
-    shows "(\<lambda>z. if z = a then deriv g a
-                 else f z - g a/(z - a)) holomorphic_on s"
-  using pole_lemma_open [OF holg s]
-  by (rule holomorphic_transform) (auto simp: eq divide_simps)
-
-proposition pole_theorem_0:
-  assumes holg: "g holomorphic_on s" and a: "a \<in> interior s"
-      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
-      and [simp]: "f a = deriv g a" "g a = 0"
-    shows "f holomorphic_on s"
-  using pole_theorem [OF holg a eq]
-  by (rule holomorphic_transform) (auto simp: eq divide_simps)
-
-proposition pole_theorem_open_0:
-  assumes holg: "g holomorphic_on s" and s: "open s"
-      and eq: "\<And>z. z \<in> s - {a} \<Longrightarrow> g z = (z - a) * f z"
-      and [simp]: "f a = deriv g a" "g a = 0"
-    shows "f holomorphic_on s"
-  using pole_theorem_open [OF holg s eq]
-  by (rule holomorphic_transform) (auto simp: eq divide_simps)
-
-lemma pole_theorem_analytic:
-  assumes g: "g analytic_on s"
-      and eq: "\<And>z. z \<in> s
-             \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
-    shows "(\<lambda>z. if z = a then deriv g a
-                 else f z - g a/(z - a)) analytic_on s"
-using g
-apply (simp add: analytic_on_def Ball_def)
-apply (safe elim!: all_forward dest!: eq)
-apply (rule_tac x="min d e" in exI, simp)
-apply (rule pole_theorem_open)
-apply (auto simp: holomorphic_on_subset subset_ball)
-done
-
-lemma pole_theorem_analytic_0:
-  assumes g: "g analytic_on s"
-      and eq: "\<And>z. z \<in> s \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>w \<in> ball z d - {a}. g w = (w - a) * f w)"
-      and [simp]: "f a = deriv g a" "g a = 0"
-    shows "f analytic_on s"
-proof -
-  have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
-    by auto
-  show ?thesis
-    using pole_theorem_analytic [OF g eq] by simp
-qed
-
-lemma pole_theorem_analytic_open_superset:
-  assumes g: "g analytic_on s" and "s \<subseteq> t" "open t"
-      and eq: "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
-    shows "(\<lambda>z. if z = a then deriv g a
-                 else f z - g a/(z - a)) analytic_on s"
-  apply (rule pole_theorem_analytic [OF g])
-  apply (rule openE [OF \<open>open t\<close>])
-  using assms eq by auto
-
-lemma pole_theorem_analytic_open_superset_0:
-  assumes g: "g analytic_on s" "s \<subseteq> t" "open t" "\<And>z. z \<in> t - {a} \<Longrightarrow> g z = (z - a) * f z"
-      and [simp]: "f a = deriv g a" "g a = 0"
-    shows "f analytic_on s"
-proof -
-  have [simp]: "(\<lambda>z. if z = a then deriv g a else f z - g a / (z - a)) = f"
-    by auto
-  have "(\<lambda>z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on s"
-    by (rule pole_theorem_analytic_open_superset [OF g])
-  then show ?thesis by simp
-qed
-
-
-
-subsection\<open>General, homology form of Cauchy's theorem.\<close>
-
-text\<open>Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\<close>
-
-lemma contour_integral_continuous_on_linepath_2D:
-  assumes "open u" and cont_dw: "\<And>w. w \<in> u \<Longrightarrow> F w contour_integrable_on (linepath a b)"
-      and cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). F x y)"
-      and abu: "closed_segment a b \<subseteq> u"
-    shows "continuous_on u (\<lambda>w. contour_integral (linepath a b) (F w))"
-proof -
-  have *: "\<exists>d>0. \<forall>x'\<in>u. dist x' w < d \<longrightarrow>
-                         dist (contour_integral (linepath a b) (F x'))
-                              (contour_integral (linepath a b) (F w)) \<le> \<epsilon>"
-          if "w \<in> u" "0 < \<epsilon>" "a \<noteq> b" for w \<epsilon>
-  proof -
-    obtain \<delta> where "\<delta>>0" and \<delta>: "cball w \<delta> \<subseteq> u" using open_contains_cball \<open>open u\<close> \<open>w \<in> u\<close> by force
-    let ?TZ = "{(t,z) |t z. t \<in> cball w \<delta> \<and> z \<in> closed_segment a b}"
-    have "uniformly_continuous_on ?TZ (\<lambda>(x,y). F x y)"
-      apply (rule compact_uniformly_continuous)
-      apply (rule continuous_on_subset[OF cond_uu])
-      using abu \<delta>
-      apply (auto simp: compact_Times simp del: mem_cball)
-      done
-    then obtain \<eta> where "\<eta>>0"
-        and \<eta>: "\<And>x x'. \<lbrakk>x\<in>?TZ; x'\<in>?TZ; dist x' x < \<eta>\<rbrakk> \<Longrightarrow>
-                         dist ((\<lambda>(x,y). F x y) x') ((\<lambda>(x,y). F x y) x) < \<epsilon>/norm(b - a)"
-      apply (rule uniformly_continuous_onE [where e = "\<epsilon>/norm(b - a)"])
-      using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> by auto
-    have \<eta>: "\<lbrakk>norm (w - x1) \<le> \<delta>;   x2 \<in> closed_segment a b;
-              norm (w - x1') \<le> \<delta>;  x2' \<in> closed_segment a b; norm ((x1', x2') - (x1, x2)) < \<eta>\<rbrakk>
-              \<Longrightarrow> norm (F x1' x2' - F x1 x2) \<le> \<epsilon> / cmod (b - a)"
-             for x1 x2 x1' x2'
-      using \<eta> [of "(x1,x2)" "(x1',x2')"] by (force simp add: dist_norm)
-    have le_ee: "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>"
-                if "x' \<in> u" "cmod (x' - w) < \<delta>" "cmod (x' - w) < \<eta>"  for x'
-    proof -
-      have "cmod (contour_integral (linepath a b) (\<lambda>x. F x' x - F w x)) \<le> \<epsilon>/norm(b - a) * norm(b - a)"
-        apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \<eta>])
-        apply (rule contour_integrable_diff [OF cont_dw cont_dw])
-        using \<open>0 < \<epsilon>\<close> \<open>a \<noteq> b\<close> \<open>0 < \<delta>\<close> \<open>w \<in> u\<close> that
-        apply (auto simp: norm_minus_commute)
-        done
-      also have "... = \<epsilon>" using \<open>a \<noteq> b\<close> by simp
-      finally show ?thesis .
-    qed
-    show ?thesis
-      apply (rule_tac x="min \<delta> \<eta>" in exI)
-      using \<open>0 < \<delta>\<close> \<open>0 < \<eta>\<close>
-      apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \<open>w \<in> u\<close> intro: le_ee)
-      done
-  qed
-  show ?thesis
-    apply (rule continuous_onI)
-    apply (cases "a=b")
-    apply (auto intro: *)
-    done
-qed
-
-text\<open>This version has @{term"polynomial_function \<gamma>"} as an additional assumption.\<close>
-lemma Cauchy_integral_formula_global_weak:
-    assumes u: "open u" and holf: "f holomorphic_on u"
-        and z: "z \<in> u" and \<gamma>: "polynomial_function \<gamma>"
-        and pasz: "path_image \<gamma> \<subseteq> u - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-        and zero: "\<And>w. w \<notin> u \<Longrightarrow> winding_number \<gamma> w = 0"
-      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
-proof -
-  obtain \<gamma>' where pf\<gamma>': "polynomial_function \<gamma>'" and \<gamma>': "\<And>x. (\<gamma> has_vector_derivative (\<gamma>' x)) (at x)"
-    using has_vector_derivative_polynomial_function [OF \<gamma>] by blast
-  then have "bounded(path_image \<gamma>')"
-    by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function)
-  then obtain B where "B>0" and B: "\<And>x. x \<in> path_image \<gamma>' \<Longrightarrow> norm x \<le> B"
-    using bounded_pos by force
-  define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w
-  define v where "v = {w. w \<notin> path_image \<gamma> \<and> winding_number \<gamma> w = 0}"
-  have "path \<gamma>" "valid_path \<gamma>" using \<gamma>
-    by (auto simp: path_polynomial_function valid_path_polynomial_function)
-  then have ov: "open v"
-    by (simp add: v_def open_winding_number_levelsets loop)
-  have uv_Un: "u \<union> v = UNIV"
-    using pasz zero by (auto simp: v_def)
-  have conf: "continuous_on u f"
-    by (metis holf holomorphic_on_imp_continuous_on)
-  have hol_d: "(d y) holomorphic_on u" if "y \<in> u" for y
-  proof -
-    have *: "(\<lambda>c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on u"
-      by (simp add: holf pole_lemma_open u)
-    then have "isCont (\<lambda>x. if x = y then deriv f y else (f x - f y) / (x - y)) y"
-      using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that u by fastforce
-    then have "continuous_on u (d y)"
-      apply (simp add: d_def continuous_on_eq_continuous_at u, clarify)
-      using * holomorphic_on_def
-      by (meson field_differentiable_within_open field_differentiable_imp_continuous_at u)
-    moreover have "d y holomorphic_on u - {y}"
-      apply (simp add: d_def holomorphic_on_open u open_delete field_differentiable_def [symmetric])
-      apply (intro ballI)
-      apply (rename_tac w)
-      apply (rule_tac d="dist w y" and f = "\<lambda>w. (f w - f y)/(w - y)" in field_differentiable_transform_within)
-      apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros)
-      using analytic_on_imp_differentiable_at analytic_on_open holf u apply blast
-      done
-    ultimately show ?thesis
-      by (rule no_isolated_singularity) (auto simp: u)
-  qed
-  have cint_fxy: "(\<lambda>x. (f x - f y) / (x - y)) contour_integrable_on \<gamma>" if "y \<notin> path_image \<gamma>" for y
-    apply (rule contour_integrable_holomorphic_simple [where s = "u-{y}"])
-    using \<open>valid_path \<gamma>\<close> pasz
-    apply (auto simp: u open_delete)
-    apply (rule continuous_intros holomorphic_intros continuous_on_subset [OF conf] holomorphic_on_subset [OF holf] |
-                force simp add: that)+
-    done
-  define h where
-    "h z = (if z \<in> u then contour_integral \<gamma> (d z) else contour_integral \<gamma> (\<lambda>w. f w/(w - z)))" for z
-  have U: "\<And>z. z \<in> u \<Longrightarrow> ((d z) has_contour_integral h z) \<gamma>"
-    apply (simp add: h_def)
-    apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
-    using u pasz \<open>valid_path \<gamma>\<close>
-    apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
-    done
-  have V: "((\<lambda>w. f w / (w - z)) has_contour_integral h z) \<gamma>" if z: "z \<in> v" for z
-  proof -
-    have 0: "0 = (f z) * 2 * of_real (2 * pi) * \<i> * winding_number \<gamma> z"
-      using v_def z by auto
-    then have "((\<lambda>x. 1 / (x - z)) has_contour_integral 0) \<gamma>"
-     using z v_def  has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close>] by fastforce
-    then have "((\<lambda>x. f z * (1 / (x - z))) has_contour_integral 0) \<gamma>"
-      using has_contour_integral_lmul by fastforce
-    then have "((\<lambda>x. f z / (x - z)) has_contour_integral 0) \<gamma>"
-      by (simp add: divide_simps)
-    moreover have "((\<lambda>x. (f x - f z) / (x - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
-      using z
-      apply (auto simp: v_def)
-      apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy)
-      done
-    ultimately have *: "((\<lambda>x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \<gamma> (d z))) \<gamma>"
-      by (rule has_contour_integral_add)
-    have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (d z)) \<gamma>"
-            if  "z \<in> u"
-      using * by (auto simp: divide_simps has_contour_integral_eq)
-    moreover have "((\<lambda>w. f w / (w - z)) has_contour_integral contour_integral \<gamma> (\<lambda>w. f w / (w - z))) \<gamma>"
-            if "z \<notin> u"
-      apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where s=u]])
-      using u pasz \<open>valid_path \<gamma>\<close> that
-      apply (auto intro: holomorphic_on_imp_continuous_on hol_d)
-      apply (rule continuous_intros conf holomorphic_intros holf | force)+
-      done
-    ultimately show ?thesis
-      using z by (simp add: h_def)
-  qed
-  have znot: "z \<notin> path_image \<gamma>"
-    using pasz by blast
-  obtain d0 where "d0>0" and d0: "\<And>x y. x \<in> path_image \<gamma> \<Longrightarrow> y \<in> - u \<Longrightarrow> d0 \<le> dist x y"
-    using separate_compact_closed [of "path_image \<gamma>" "-u"] pasz u
-    by (fastforce simp add: \<open>path \<gamma>\<close> compact_path_image)
-  obtain dd where "0 < dd" and dd: "{y + k | y k. y \<in> path_image \<gamma> \<and> k \<in> ball 0 dd} \<subseteq> u"
-    apply (rule that [of "d0/2"])
-    using \<open>0 < d0\<close>
-    apply (auto simp: dist_norm dest: d0)
-    done
-  have "\<And>x x'. \<lbrakk>x \<in> path_image \<gamma>; dist x x' * 2 < dd\<rbrakk> \<Longrightarrow> \<exists>y k. x' = y + k \<and> y \<in> path_image \<gamma> \<and> dist 0 k * 2 \<le> dd"
-    apply (rule_tac x=x in exI)
-    apply (rule_tac x="x'-x" in exI)
-    apply (force simp add: dist_norm)
-    done
-  then have 1: "path_image \<gamma> \<subseteq> interior {y + k |y k. y \<in> path_image \<gamma> \<and> k \<in> cball 0 (dd / 2)}"
-    apply (clarsimp simp add: mem_interior)
-    using \<open>0 < dd\<close>
-    apply (rule_tac x="dd/2" in exI, auto)
-    done
-  obtain t where "compact t" and subt: "path_image \<gamma> \<subseteq> interior t" and t: "t \<subseteq> u"
-    apply (rule that [OF _ 1])
-    apply (fastforce simp add: \<open>valid_path \<gamma>\<close> compact_valid_path_image intro!: compact_sums)
-    apply (rule order_trans [OF _ dd])
-    using \<open>0 < dd\<close> by fastforce
-  obtain L where "L>0"
-           and L: "\<And>f B. \<lbrakk>f holomorphic_on interior t; \<And>z. z\<in>interior t \<Longrightarrow> cmod (f z) \<le> B\<rbrakk> \<Longrightarrow>
-                         cmod (contour_integral \<gamma> f) \<le> L * B"
-      using contour_integral_bound_exists [OF open_interior \<open>valid_path \<gamma>\<close> subt]
-      by blast
-  have "bounded(f ` t)"
-    by (meson \<open>compact t\<close> compact_continuous_image compact_imp_bounded conf continuous_on_subset t)
-  then obtain D where "D>0" and D: "\<And>x. x \<in> t \<Longrightarrow> norm (f x) \<le> D"
-    by (auto simp: bounded_pos)
-  obtain C where "C>0" and C: "\<And>x. x \<in> t \<Longrightarrow> norm x \<le> C"
-    using \<open>compact t\<close> bounded_pos compact_imp_bounded by force
-  have "dist (h y) 0 \<le> e" if "0 < e" and le: "D * L / e + C \<le> cmod y" for e y
-  proof -
-    have "D * L / e > 0"  using \<open>D>0\<close> \<open>L>0\<close> \<open>e>0\<close> by simp
-    with le have ybig: "norm y > C" by force
-    with C have "y \<notin> t"  by force
-    then have ynot: "y \<notin> path_image \<gamma>"
-      using subt interior_subset by blast
-    have [simp]: "winding_number \<gamma> y = 0"
-      apply (rule winding_number_zero_outside [of _ "cball 0 C"])
-      using ybig interior_subset subt
-      apply (force simp add: loop \<open>path \<gamma>\<close> dist_norm intro!: C)+
-      done
-    have [simp]: "h y = contour_integral \<gamma> (\<lambda>w. f w/(w - y))"
-      by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V)
-    have holint: "(\<lambda>w. f w / (w - y)) holomorphic_on interior t"
-      apply (rule holomorphic_on_divide)
-      using holf holomorphic_on_subset interior_subset t apply blast
-      apply (rule holomorphic_intros)+
-      using \<open>y \<notin> t\<close> interior_subset by auto
-    have leD: "cmod (f z / (z - y)) \<le> D * (e / L / D)" if z: "z \<in> interior t" for z
-    proof -
-      have "D * L / e + cmod z \<le> cmod y"
-        using le C [of z] z using interior_subset by force
-      then have DL2: "D * L / e \<le> cmod (z - y)"
-        using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute)
-      have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))"
-        by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse)
-      also have "... \<le> D * (e / L / D)"
-        apply (rule mult_mono)
-        using that D interior_subset apply blast
-        using \<open>L>0\<close> \<open>e>0\<close> \<open>D>0\<close> DL2
-        apply (auto simp: norm_divide divide_simps algebra_simps)
-        done
-      finally show ?thesis .
-    qed
-    have "dist (h y) 0 = cmod (contour_integral \<gamma> (\<lambda>w. f w / (w - y)))"
-      by (simp add: dist_norm)
-    also have "... \<le> L * (D * (e / L / D))"
-      by (rule L [OF holint leD])
-    also have "... = e"
-      using  \<open>L>0\<close> \<open>0 < D\<close> by auto
-    finally show ?thesis .
-  qed
-  then have "(h \<longlongrightarrow> 0) at_infinity"
-    by (meson Lim_at_infinityI)
-  moreover have "h holomorphic_on UNIV"
-  proof -
-    have con_ff: "continuous (at (x,z)) (\<lambda>(x,y). (f y - f x) / (y - x))"
-                 if "x \<in> u" "z \<in> u" "x \<noteq> z" for x z
-      using that conf
-      apply (simp add: split_def continuous_on_eq_continuous_at u)
-      apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+
-      done
-    have con_fstsnd: "continuous_on UNIV (\<lambda>x. (fst x - snd x) ::complex)"
-      by (rule continuous_intros)+
-    have open_uu_Id: "open (u \<times> u - Id)"
-      apply (rule open_Diff)
-      apply (simp add: open_Times u)
-      using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0]
-      apply (auto simp: Id_fstsnd_eq algebra_simps)
-      done
-    have con_derf: "continuous (at z) (deriv f)" if "z \<in> u" for z
-      apply (rule continuous_on_interior [of u])
-      apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on u)
-      by (simp add: interior_open that u)
-    have tendsto_f': "((\<lambda>(x,y). if y = x then deriv f (x)
-                                else (f (y) - f (x)) / (y - x)) \<longlongrightarrow> deriv f x)
-                      (at (x, x) within u \<times> u)" if "x \<in> u" for x
-    proof (rule Lim_withinI)
-      fix e::real assume "0 < e"
-      obtain k1 where "k1>0" and k1: "\<And>x'. norm (x' - x) \<le> k1 \<Longrightarrow> norm (deriv f x' - deriv f x) < e"
-        using \<open>0 < e\<close> continuous_within_E [OF con_derf [OF \<open>x \<in> u\<close>]]
-        by (metis UNIV_I dist_norm)
-      obtain k2 where "k2>0" and k2: "ball x k2 \<subseteq> u" by (blast intro: openE [OF u] \<open>x \<in> u\<close>)
-      have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e"
-                    if "z' \<noteq> x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2"
-                 for x' z'
-      proof -
-        have cs_less: "w \<in> closed_segment x' z' \<Longrightarrow> cmod (w - x) \<le> norm (x'-x, z'-x)" for w
-          apply (drule segment_furthest_le [where y=x])
-          by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans)
-        have derf_le: "w \<in> closed_segment x' z' \<Longrightarrow> z' \<noteq> x' \<Longrightarrow> cmod (deriv f w - deriv f x) \<le> e" for w
-          by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans)
-        have f_has_der: "\<And>x. x \<in> u \<Longrightarrow> (f has_field_derivative deriv f x) (at x within u)"
-          by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def u)
-        have "closed_segment x' z' \<subseteq> u"
-          by (rule order_trans [OF _ k2]) (simp add: cs_less  le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff)
-        then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')"
-          using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz  by simp
-        then have *: "((\<lambda>x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')"
-          by (rule has_contour_integral_div)
-        have "norm ((f z' - f x') / (z' - x') - deriv f x) \<le> e/norm(z' - x') * norm(z' - x')"
-          apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]])
-          using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']]
-                 \<open>e > 0\<close>  \<open>z' \<noteq> x'\<close>
-          apply (auto simp: norm_divide divide_simps derf_le)
-          done
-        also have "... \<le> e" using \<open>0 < e\<close> by simp
-        finally show ?thesis .
-      qed
-      show "\<exists>d>0. \<forall>xa\<in>u \<times> u.
-                  0 < dist xa (x, x) \<and> dist xa (x, x) < d \<longrightarrow>
-                  dist (case xa of (x, y) \<Rightarrow> if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \<le> e"
-        apply (rule_tac x="min k1 k2" in exI)
-        using \<open>k1>0\<close> \<open>k2>0\<close> \<open>e>0\<close>
-        apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le)
-        done
-    qed
-    have con_pa_f: "continuous_on (path_image \<gamma>) f"
-      by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt t)
-    have le_B: "\<And>t. t \<in> {0..1} \<Longrightarrow> cmod (vector_derivative \<gamma> (at t)) \<le> B"
-      apply (rule B)
-      using \<gamma>' using path_image_def vector_derivative_at by fastforce
-    have f_has_cint: "\<And>w. w \<in> v - path_image \<gamma> \<Longrightarrow> ((\<lambda>u. f u / (u - w) ^ 1) has_contour_integral h w) \<gamma>"
-      by (simp add: V)
-    have cond_uu: "continuous_on (u \<times> u) (\<lambda>(x,y). d x y)"
-      apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f')
-      apply (simp add: tendsto_within_open_NO_MATCH open_Times u, clarify)
-      apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\<lambda>(x,y). (f y - f x) / (y - x))"])
-      using con_ff
-      apply (auto simp: continuous_within)
-      done
-    have hol_dw: "(\<lambda>z. d z w) holomorphic_on u" if "w \<in> u" for w
-    proof -
-      have "continuous_on u ((\<lambda>(x,y). d x y) o (\<lambda>z. (w,z)))"
-        by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+
-      then have *: "continuous_on u (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z))"
-        by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps)
-      have **: "\<And>x. \<lbrakk>x \<in> u; x \<noteq> w\<rbrakk> \<Longrightarrow> (\<lambda>z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x"
-        apply (rule_tac f = "\<lambda>x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within)
-        apply (rule u derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp add: dist_commute)+
-        done
-      show ?thesis
-        unfolding d_def
-        apply (rule no_isolated_singularity [OF * _ u, where k = "{w}"])
-        apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff u **)
-        done
-    qed
-    { fix a b
-      assume abu: "closed_segment a b \<subseteq> u"
-      then have "\<And>w. w \<in> u \<Longrightarrow> (\<lambda>z. d z w) contour_integrable_on (linepath a b)"
-        by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on)
-      then have cont_cint_d: "continuous_on u (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
-        apply (rule contour_integral_continuous_on_linepath_2D [OF \<open>open u\<close> _ _ abu])
-        apply (auto simp: intro: continuous_on_swap_args cond_uu)
-        done
-      have cont_cint_d\<gamma>: "continuous_on {0..1} ((\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) o \<gamma>)"
-        apply (rule continuous_on_compose)
-        using \<open>path \<gamma>\<close> path_def pasz
-        apply (auto intro!: continuous_on_subset [OF cont_cint_d])
-        apply (force simp add: path_image_def)
-        done
-      have cint_cint: "(\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w)) contour_integrable_on \<gamma>"
-        apply (simp add: contour_integrable_on)
-        apply (rule integrable_continuous_real)
-        apply (rule continuous_on_mult [OF cont_cint_d\<gamma> [unfolded o_def]])
-        using pf\<gamma>'
-        by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \<gamma>'])
-      have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\<lambda>z. contour_integral \<gamma> (d z))"
-        using abu  by (force simp add: h_def intro: contour_integral_eq)
-      also have "... =  contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))"
-        apply (rule contour_integral_swap)
-        apply (rule continuous_on_subset [OF cond_uu])
-        using abu pasz \<open>valid_path \<gamma>\<close>
-        apply (auto intro!: continuous_intros)
-        by (metis \<gamma>' continuous_on_eq path_def path_polynomial_function pf\<gamma>' vector_derivative_at)
-      finally have cint_h_eq:
-          "contour_integral (linepath a b) h =
-                    contour_integral \<gamma> (\<lambda>w. contour_integral (linepath a b) (\<lambda>z. d z w))" .
-      note cint_cint cint_h_eq
-    } note cint_h = this
-    have conthu: "continuous_on u h"
-    proof (simp add: continuous_on_sequentially, clarify)
-      fix a x
-      assume x: "x \<in> u" and au: "\<forall>n. a n \<in> u" and ax: "a \<longlonglongrightarrow> x"
-      then have A1: "\<forall>\<^sub>F n in sequentially. d (a n) contour_integrable_on \<gamma>"
-        by (meson U contour_integrable_on_def eventuallyI)
-      obtain dd where "dd>0" and dd: "cball x dd \<subseteq> u" using open_contains_cball u x by force
-      have A2: "\<forall>\<^sub>F n in sequentially. \<forall>xa\<in>path_image \<gamma>. cmod (d (a n) xa - d x xa) < ee" if "0 < ee" for ee
-      proof -
-        let ?ddpa = "{(w,z) |w z. w \<in> cball x dd \<and> z \<in> path_image \<gamma>}"
-        have "uniformly_continuous_on ?ddpa (\<lambda>(x,y). d x y)"
-          apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]])
-          using dd pasz \<open>valid_path \<gamma>\<close>
-          apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball)
-          done
-        then obtain kk where "kk>0"
-            and kk: "\<And>x x'. \<lbrakk>x\<in>?ddpa; x'\<in>?ddpa; dist x' x < kk\<rbrakk> \<Longrightarrow>
-                             dist ((\<lambda>(x,y). d x y) x') ((\<lambda>(x,y). d x y) x) < ee"
-          apply (rule uniformly_continuous_onE [where e = ee])
-          using \<open>0 < ee\<close> by auto
-        have kk: "\<lbrakk>norm (w - x) \<le> dd; z \<in> path_image \<gamma>; norm ((w, z) - (x, z)) < kk\<rbrakk> \<Longrightarrow> norm (d w z - d x z) < ee"
-                 for  w z
-          using \<open>dd>0\<close> kk [of "(x,z)" "(w,z)"] by (force simp add: norm_minus_commute dist_norm)
-        show ?thesis
-          using ax unfolding lim_sequentially eventually_sequentially
-          apply (drule_tac x="min dd kk" in spec)
-          using \<open>dd > 0\<close> \<open>kk > 0\<close>
-          apply (fastforce simp: kk dist_norm)
-          done
-      qed
-      have tendsto_hx: "(\<lambda>n. contour_integral \<gamma> (d (a n))) \<longlonglongrightarrow> h x"
-        apply (simp add: contour_integral_unique [OF U, symmetric] x)
-        apply (rule contour_integral_uniform_limit [OF A1 A2 le_B])
-        apply (auto simp: \<open>valid_path \<gamma>\<close>)
-        done
-      then show "(h \<circ> a) \<longlonglongrightarrow> h x"
-        by (simp add: h_def x au o_def)
-    qed
-    show ?thesis
-    proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify)
-      fix z0
-      consider "z0 \<in> v" | "z0 \<in> u" using uv_Un by blast
-      then show "h field_differentiable at z0"
-      proof cases
-        assume "z0 \<in> v" then show ?thesis
-          using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \<open>valid_path \<gamma>\<close>
-          by (auto simp: field_differentiable_def v_def)
-      next
-        assume "z0 \<in> u" then
-        obtain e where "e>0" and e: "ball z0 e \<subseteq> u" by (blast intro: openE [OF u])
-        have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0"
-                if abc_subset: "convex hull {a, b, c} \<subseteq> ball z0 e"  for a b c
-        proof -
-          have *: "\<And>x1 x2 z. z \<in> u \<Longrightarrow> closed_segment x1 x2 \<subseteq> u \<Longrightarrow> (\<lambda>w. d w z) contour_integrable_on linepath x1 x2"
-            using  hol_dw holomorphic_on_imp_continuous_on u
-            by (auto intro!: contour_integrable_holomorphic_simple)
-          have abc: "closed_segment a b \<subseteq> u"  "closed_segment b c \<subseteq> u"  "closed_segment c a \<subseteq> u"
-            using that e segments_subset_convex_hull by fastforce+
-          have eq0: "\<And>w. w \<in> u \<Longrightarrow> contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\<lambda>z. d z w) = 0"
-            apply (rule contour_integral_unique [OF Cauchy_theorem_triangle])
-            apply (rule holomorphic_on_subset [OF hol_dw])
-            using e abc_subset by auto
-          have "contour_integral \<gamma>
-                   (\<lambda>x. contour_integral (linepath a b) (\<lambda>z. d z x) +
-                        (contour_integral (linepath b c) (\<lambda>z. d z x) +
-                         contour_integral (linepath c a) (\<lambda>z. d z x)))  =  0"
-            apply (rule contour_integral_eq_0)
-            using abc pasz u
-            apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+
-            done
-          then show ?thesis
-            by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac)
-        qed
-        show ?thesis
-          using e \<open>e > 0\<close>
-          by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic
-                           Morera_triangle continuous_on_subset [OF conthu] *)
-      qed
-    qed
-  qed
-  ultimately have [simp]: "h z = 0" for z
-    by (meson Liouville_weak)
-  have "((\<lambda>w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z) \<gamma>"
-    by (rule has_contour_integral_winding_number [OF \<open>valid_path \<gamma>\<close> znot])
-  then have "((\<lambda>w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
-    by (metis mult.commute has_contour_integral_lmul)
-  then have 1: "((\<lambda>w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \<i> * winding_number \<gamma> z * f z) \<gamma>"
-    by (simp add: divide_simps)
-  moreover have 2: "((\<lambda>w. (f w - f z) / (w - z)) has_contour_integral 0) \<gamma>"
-    using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\<lambda>w. (f w - f z)/(w - z)"])
-  show ?thesis
-    using has_contour_integral_add [OF 1 2]  by (simp add: diff_divide_distrib)
-qed
-
-
-theorem Cauchy_integral_formula_global:
-    assumes s: "open s" and holf: "f holomorphic_on s"
-        and z: "z \<in> s" and vpg: "valid_path \<gamma>"
-        and pasz: "path_image \<gamma> \<subseteq> s - {z}" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-        and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
-      shows "((\<lambda>w. f w / (w - z)) has_contour_integral (2*pi * \<i> * winding_number \<gamma> z * f z)) \<gamma>"
-proof -
-  have "path \<gamma>" using vpg by (blast intro: valid_path_imp_path)
-  have hols: "(\<lambda>w. f w / (w - z)) holomorphic_on s - {z}" "(\<lambda>w. 1 / (w - z)) holomorphic_on s - {z}"
-    by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+
-  then have cint_fw: "(\<lambda>w. f w / (w - z)) contour_integrable_on \<gamma>"
-    by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete s vpg pasz)
-  obtain d where "d>0"
-      and d: "\<And>g h. \<lbrakk>valid_path g; valid_path h; \<forall>t\<in>{0..1}. cmod (g t - \<gamma> t) < d \<and> cmod (h t - \<gamma> t) < d;
-                     pathstart h = pathstart g \<and> pathfinish h = pathfinish g\<rbrakk>
-                     \<Longrightarrow> path_image h \<subseteq> s - {z} \<and> (\<forall>f. f holomorphic_on s - {z} \<longrightarrow> contour_integral h f = contour_integral g f)"
-    using contour_integral_nearby_ends [OF _ \<open>path \<gamma>\<close> pasz] s by (simp add: open_Diff) metis
-  obtain p where polyp: "polynomial_function p"
-             and ps: "pathstart p = pathstart \<gamma>" and pf: "pathfinish p = pathfinish \<gamma>" and led: "\<forall>t\<in>{0..1}. cmod (p t - \<gamma> t) < d"
-    using path_approx_polynomial_function [OF \<open>path \<gamma>\<close> \<open>d > 0\<close>] by blast
-  then have ploop: "pathfinish p = pathstart p" using loop by auto
-  have vpp: "valid_path p"  using polyp valid_path_polynomial_function by blast
-  have [simp]: "z \<notin> path_image \<gamma>" using pasz by blast
-  have paps: "path_image p \<subseteq> s - {z}" and cint_eq: "(\<And>f. f holomorphic_on s - {z} \<Longrightarrow> contour_integral p f = contour_integral \<gamma> f)"
-    using pf ps led d [OF vpg vpp] \<open>d > 0\<close> by auto
-  have wn_eq: "winding_number p z = winding_number \<gamma> z"
-    using vpp paps
-    by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols)
-  have "winding_number p w = winding_number \<gamma> w" if "w \<notin> s" for w
-  proof -
-    have hol: "(\<lambda>v. 1 / (v - w)) holomorphic_on s - {z}"
-      using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf])
-   have "w \<notin> path_image p" "w \<notin> path_image \<gamma>" using paps pasz that by auto
-   then show ?thesis
-    using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol])
-  qed
-  then have wn0: "\<And>w. w \<notin> s \<Longrightarrow> winding_number p w = 0"
-    by (simp add: zero)
-  show ?thesis
-    using Cauchy_integral_formula_global_weak [OF s holf z polyp paps ploop wn0] hols
-    by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq)
-qed
-
-theorem Cauchy_theorem_global:
-    assumes s: "open s" and holf: "f holomorphic_on s"
-        and vpg: "valid_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>"
-        and pas: "path_image \<gamma> \<subseteq> s"
-        and zero: "\<And>w. w \<notin> s \<Longrightarrow> winding_number \<gamma> w = 0"
-      shows "(f has_contour_integral 0) \<gamma>"
-proof -
-  obtain z where "z \<in> s" and znot: "z \<notin> path_image \<gamma>"
-  proof -
-    have "compact (path_image \<gamma>)"
-      using compact_valid_path_image vpg by blast
-    then have "path_image \<gamma> \<noteq> s"
-      by (metis (no_types) compact_open path_image_nonempty s)
-    with pas show ?thesis by (blast intro: that)
-  qed
-  then have pasz: "path_image \<gamma> \<subseteq> s - {z}" using pas by blast
-  have hol: "(\<lambda>w. (w - z) * f w) holomorphic_on s"
-    by (rule holomorphic_intros holf)+
-  show ?thesis
-    using Cauchy_integral_formula_global [OF s hol \<open>z \<in> s\<close> vpg pasz loop zero]
-    by (auto simp: znot elim!: has_contour_integral_eq)
-qed
-
-corollary Cauchy_theorem_global_outside:
-    assumes "open s" "f holomorphic_on s" "valid_path \<gamma>"  "pathfinish \<gamma> = pathstart \<gamma>" "path_image \<gamma> \<subseteq> s"
-            "\<And>w. w \<notin> s \<Longrightarrow> w \<in> outside(path_image \<gamma>)"
-      shows "(f has_contour_integral 0) \<gamma>"
-by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path)
-
-
-end
--- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -5,7 +5,7 @@
 theory Complex_Transcendental
 imports
   Complex_Analysis_Basics
-  Summation
+  Summation_Tests
 begin
 
 (* TODO: Figure out what to do with Möbius transformations *)
@@ -1647,7 +1647,7 @@
 
 lemma norm_powr_real_powr:
   "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
-  by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 
+  by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0
                                      complex_is_Real_iff in_Reals_norm complex_eq_iff)
 
 lemma tendsto_ln_complex [tendsto_intros]:
--- a/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -5,7 +5,7 @@
 text\<open>Also Cauchy's residue theorem by Wenda Li (2016)\<close>
 
 theory Conformal_Mappings
-imports "~~/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm"
+imports "~~/src/HOL/Multivariate_Analysis/Cauchy_Integral_Theorem"
 
 begin
 
@@ -1897,7 +1897,7 @@
   qed
 qed
 
-lemma Bloch_lemma: 
+lemma Bloch_lemma:
   assumes holf: "f holomorphic_on cball a r" and "0 < r"
       and le: "\<And>z. z \<in> ball a r \<Longrightarrow> norm(deriv f z) \<le> 2 * norm(deriv f a)"
     shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \<subseteq> f ` ball a r"
@@ -1928,7 +1928,7 @@
     done
 qed
 
-proposition Bloch_unit: 
+proposition Bloch_unit:
   assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1"
   obtains b r where "1/12 < r" "ball b r \<subseteq> f ` (ball a 1)"
 proof -
@@ -1947,23 +1947,23 @@
     by (rule compact_continuous_image [OF _ compact_cball])
   have 2: "g ` cball a r \<noteq> {}"
     using \<open>r > 0\<close> by auto
-  obtain p where pr: "p \<in> cball a r" 
+  obtain p where pr: "p \<in> cball a r"
              and pge: "\<And>y. y \<in> cball a r \<Longrightarrow> norm (g y) \<le> norm (g p)"
     using distance_attains_sup [OF 1 2, of 0] by force
   define t where "t = (r - norm(p - a)) / 2"
   have "norm (p - a) \<noteq> r"
     using pge [of a] \<open>r > 0\<close> by (auto simp: g_def norm_mult)
-  then have "norm (p - a) < r" using pr 
+  then have "norm (p - a) < r" using pr
     by (simp add: norm_minus_commute dist_norm)
-  then have "0 < t" 
+  then have "0 < t"
     by (simp add: t_def)
   have cpt: "cball p t \<subseteq> ball a r"
     using \<open>0 < t\<close> by (simp add: cball_subset_ball_iff dist_norm t_def field_simps)
-  have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" 
+  have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
             if "y \<in> cball a r" for y
   proof -
     have [simp]: "norm (y - a) \<le> r"
-      using that by (simp add: dist_norm norm_minus_commute) 
+      using that by (simp add: dist_norm norm_minus_commute)
     have "norm (g y) \<le> norm (g p)"
       using pge [OF that] by simp
     then have "norm (deriv f y) * abs (r - norm (y - a)) \<le> norm (deriv f p) * abs (r - norm (p - a))"
@@ -1982,14 +1982,14 @@
       by (meson ball_subset_cball subsetD cpt that)
     then have "norm(z - a) < r"
       by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that)
-    have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)" 
+    have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \<le> norm (deriv f p)"
       using gen_le_dfp [OF z] by simp
-    with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close> 
+    with \<open>norm (z - a) < r\<close> \<open>norm (p - a) < r\<close>
     have "norm (deriv f z) \<le> (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)"
        by (simp add: field_simps)
     also have "... \<le> 2 * norm (deriv f p)"
       apply (rule mult_right_mono)
-      using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close> 
+      using that \<open>norm (p - a) < r\<close> \<open>norm(z - a) < r\<close>
       apply (simp_all add: field_simps t_def dist_norm [symmetric])
       using dist_triangle3 [of z a p] by linarith
     finally show ?thesis .
@@ -1998,10 +1998,10 @@
     by (rule real_less_lsqrt) (auto simp: power2_eq_square)
   then have sq3: "0 < 3 - 2 * sqrt 2" by simp
   have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r"
-    using sq3 sqrt2 by (auto simp: field_simps r_def) 
+    using sq3 sqrt2 by (auto simp: field_simps r_def)
   also have "... \<le> cmod (deriv f p) * (r - cmod (p - a))"
-    using \<open>norm (p - a) < r\<close> le_norm_dfp   by (simp add: pos_divide_le_eq)    
-  finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"  
+    using \<open>norm (p - a) < r\<close> le_norm_dfp   by (simp add: pos_divide_le_eq)
+  finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)"
     using pos_divide_less_eq half_gt_zero_iff sq3 by blast
   then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)"
     using sq3 by (simp add: mult.commute t_def)
@@ -2020,7 +2020,7 @@
 
 
 theorem Bloch:
-  assumes holf: "f holomorphic_on ball a r" and "0 < r" 
+  assumes holf: "f holomorphic_on ball a r" and "0 < r"
       and r': "r' \<le> r * norm (deriv f a) / 12"
   obtains b where "ball b r' \<subseteq> f ` (ball a r)"
 proof (cases "deriv f a = 0")
@@ -2043,7 +2043,7 @@
     apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+
     using \<open>0 < r\<close> by (simp add: C_def False)
   have "((\<lambda>z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative
-        (deriv f (a + of_real r * z) / C)) (at z)" 
+        (deriv f (a + of_real r * z) / C)) (at z)"
        if "norm z < 1" for z
   proof -
     have *: "((\<lambda>x. f (a + of_real r * x)) has_field_derivative
@@ -2065,23 +2065,23 @@
     apply (rule DERIV_chain [OF field_differentiable_derivI])
     apply (simp add: dfa)
     apply (rule derivative_eq_intros | simp add: C_def False fo)+
-    using \<open>0 < r\<close> 
+    using \<open>0 < r\<close>
     apply (simp add: C_def False fo)
     apply (simp add: derivative_intros dfa complex_derivative_chain)
     done
-  have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1 
+  have sb1: "op * (C * r) ` (\<lambda>z. f (a + of_real r * z) / (C * r)) ` ball 0 1
              \<subseteq> f ` ball a r"
     using \<open>0 < r\<close> by (auto simp: dist_norm norm_mult C_def False)
-  have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t" 
+  have sb2: "ball (C * r * b) r' \<subseteq> op * (C * r) ` ball b t"
              if "1 / 12 < t" for b t
   proof -
     have *: "r * cmod (deriv f a) / 12 \<le> r * (t * cmod (deriv f a))"
-      using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right 
+      using that \<open>0 < r\<close> less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right
       by auto
     show ?thesis
       apply clarify
       apply (rule_tac x="x / (C * r)" in image_eqI)
-      using \<open>0 < r\<close>  
+      using \<open>0 < r\<close>
       apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps)
       apply (erule less_le_trans)
       apply (rule order_trans [OF r' *])
@@ -2098,7 +2098,7 @@
 qed
 
 corollary Bloch_general:
-  assumes holf: "f holomorphic_on s" and "a \<in> s" 
+  assumes holf: "f holomorphic_on s" and "a \<in> s"
       and tle: "\<And>z. z \<in> frontier s \<Longrightarrow> t \<le> dist a z"
       and rle: "r \<le> t * norm(deriv f a) / 12"
   obtains b where "ball b r \<subseteq> f ` s"
@@ -2140,11 +2140,11 @@
     Interactive Theorem Proving\<close>
 
 definition residue :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
-  "residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e 
+  "residue f z = (SOME int. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e
     \<longrightarrow> (f has_contour_integral 2*pi* \<i> *int) (circlepath z \<epsilon>))"
 
 lemma contour_integral_circlepath_eq:
-  assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2" 
+  assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0<e1" "e1\<le>e2"
     and e2_cball:"cball z e2 \<subseteq> s"
   shows
     "f contour_integrable_on circlepath z e1"
@@ -2155,24 +2155,24 @@
   have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto
   have "e2>0" using \<open>e1>0\<close> \<open>e1\<le>e2\<close> by auto
   have zl_img:"z\<notin>path_image l"
-    proof 
+    proof
       assume "z \<in> path_image l"
       then have "e2 \<le> cmod (e2 - e1)"
         using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \<open>e1>0\<close> \<open>e2>0\<close> unfolding l_def
         by (auto simp add:closed_segment_commute)
-      thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close> 
+      thus False using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
         apply (subst (asm) norm_of_real)
         by auto
     qed
   define g where "g \<equiv> circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l"
-  show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" 
+  show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)"
     proof -
-      show "f contour_integrable_on circlepath z e2" 
-        apply (intro contour_integrable_continuous_circlepath[OF 
+      show "f contour_integrable_on circlepath z e2"
+        apply (intro contour_integrable_continuous_circlepath[OF
                 continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
         using \<open>e2>0\<close> e2_cball by auto
-      show "f contour_integrable_on (circlepath z e1)" 
-        apply (intro contour_integrable_continuous_circlepath[OF 
+      show "f contour_integrable_on (circlepath z e1)"
+        apply (intro contour_integrable_continuous_circlepath[OF
                       continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
         using \<open>e1>0\<close> \<open>e1\<le>e2\<close> e2_cball by auto
     qed
@@ -2183,7 +2183,7 @@
       hence "closed_segment (z + e2) (z + e1) \<subseteq> s - {z}" using zl_img e2_cball unfolding l_def
         by auto
       then show "f contour_integrable_on l" unfolding l_def
-        apply (intro contour_integrable_continuous_linepath[OF 
+        apply (intro contour_integrable_continuous_linepath[OF
                       continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]])
         by auto
     qed
@@ -2197,14 +2197,14 @@
       have path_img:"path_image g \<subseteq> cball z e2"
         proof -
           have "closed_segment (z + e2) (z + e1) \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1>0\<close> \<open>e1\<le>e2\<close>
-            by (intro closed_segment_subset,auto simp add:dist_norm)  
+            by (intro closed_segment_subset,auto simp add:dist_norm)
           moreover have "sphere z \<bar>e1\<bar> \<subseteq> cball z e2" using \<open>e2>0\<close> \<open>e1\<le>e2\<close> \<open>e1>0\<close> by auto
-          ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close> 
+          ultimately show ?thesis unfolding g_def l_def using \<open>e2>0\<close>
             by (simp add: path_image_join closed_segment_commute)
         qed
-      show "path_image g \<subseteq> s - {z}"  
+      show "path_image g \<subseteq> s - {z}"
         proof -
-          have "z\<notin>path_image g" using zl_img      
+          have "z\<notin>path_image g" using zl_img
             unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute)
           moreover note \<open>cball z e2 \<subseteq> s\<close> and path_img
           ultimately show ?thesis by auto
@@ -2214,12 +2214,12 @@
           have "winding_number g w = 0" when "w\<notin>s" using that e2_cball
             apply (intro winding_number_zero_outside[OF _ _ _ _ path_img])
             by (auto simp add:g_def l_def)
-          moreover have "winding_number g z=0" 
+          moreover have "winding_number g z=0"
             proof -
               let ?Wz="\<lambda>g. winding_number g z"
-              have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) 
+              have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1))
                   + ?Wz (reversepath l)"
-                using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def  
+                using \<open>e2>0\<close> \<open>e1>0\<close> zl_img unfolding g_def l_def
                 by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+
               also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))"
                 using zl_img
@@ -2237,43 +2237,43 @@
               finally show ?thesis .
             qed
           ultimately show ?thesis using that by auto
-        qed  
-    qed        
+        qed
+    qed
   then have "0 = ?ig g" using contour_integral_unique by simp
-  also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) 
+  also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1))
       + ?ig (reversepath l)"
-    unfolding g_def 
+    unfolding g_def
     by (auto simp add:contour_integrable_reversepath_eq)
   also have "... = ?ig (circlepath z e2)  - ?ig (circlepath z e1)"
     by (auto simp add:contour_integral_reversepath)
-  finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" 
+  finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f"
     by simp
 qed
 
 lemma base_residue:
-  assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})" 
+  assumes "open s" "z\<in>s" "r>0" and f_holo:"f holomorphic_on (s - {z})"
     and r_cball:"cball z r \<subseteq> s"
   shows "(f has_contour_integral 2 * pi * \<i> * (residue f z)) (circlepath z r)"
 proof -
-  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" 
+  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s"
     using open_contains_cball[of s] \<open>open s\<close> \<open>z\<in>s\<close> by auto
   define c where "c \<equiv> 2 * pi * \<i>"
   define i where "i \<equiv> contour_integral (circlepath z e) f / c"
-  have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon> 
+  have "(f has_contour_integral c*i) (circlepath z \<epsilon>)" when "\<epsilon>>0" "\<epsilon><e" for \<epsilon>
     proof -
       have "contour_integral (circlepath z e) f = contour_integral (circlepath z \<epsilon>) f"
           "f contour_integrable_on circlepath z \<epsilon>"
           "f contour_integrable_on circlepath z e"
         using \<open>\<epsilon><e\<close>
-        by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+        
+        by (intro contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> _ e_cball],auto)+
       then show ?thesis unfolding i_def c_def
         by (auto intro:has_contour_integral_integral)
     qed
-  then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)" 
+  then have "\<exists>e>0. \<forall>\<epsilon>>0. \<epsilon><e \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
     unfolding residue_def c_def
     apply (rule_tac someI[of _ i],intro  exI[where x=e])
     by (auto simp add:\<open>e>0\<close> c_def)
-  then obtain e' where "e'>0" 
+  then obtain e' where "e'>0"
       and e'_def:"\<forall>\<epsilon>>0. \<epsilon><e' \<longrightarrow> (f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
     by auto
   let ?int="\<lambda>e. contour_integral (circlepath z e) f"
@@ -2281,8 +2281,8 @@
   have "\<epsilon>>0" "\<epsilon>\<le>r" "\<epsilon><e'" using \<open>r>0\<close> \<open>e'>0\<close> unfolding \<epsilon>_def by auto
   have "(f has_contour_integral c * (residue f z)) (circlepath z \<epsilon>)"
     using e'_def[rule_format,OF \<open>\<epsilon>>0\<close> \<open>\<epsilon><e'\<close>] .
-  then show ?thesis unfolding c_def 
-    using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball] 
+  then show ?thesis unfolding c_def
+    using contour_integral_circlepath_eq[OF \<open>open s\<close> f_holo \<open>\<epsilon>>0\<close> \<open>\<epsilon>\<le>r\<close> r_cball]
     by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \<epsilon>" "circlepath z r"])
 qed
 
@@ -2292,17 +2292,17 @@
   shows "residue f z = 0"
 proof -
   define c where "c \<equiv> 2 * pi * \<i>"
-  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> 
+  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
     using open_contains_cball_eq by blast
   have "(f has_contour_integral c*residue f z) (circlepath z e)"
     using f_holo
-    by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])    
-  moreover have "(f has_contour_integral 0) (circlepath z e)" 
+    by (auto intro: base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
+  moreover have "(f has_contour_integral 0) (circlepath z e)"
     using f_holo e_cball \<open>e>0\<close>
     by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"])
   ultimately have "c*residue f z =0"
     using has_contour_integral_unique by blast
-  thus ?thesis unfolding c_def  by auto    
+  thus ?thesis unfolding c_def  by auto
 qed
 
 
@@ -2311,39 +2311,39 @@
 
 
 lemma residue_add:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       and g_holo:"g holomorphic_on s - {z}"
   shows "residue (\<lambda>z. f z + g z) z= residue f z + residue g z"
 proof -
   define c where "c \<equiv> 2 * pi * \<i>"
   define fg where "fg \<equiv> (\<lambda>z. f z+g z)"
-  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> 
+  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
     using open_contains_cball_eq by blast
-  have "(fg has_contour_integral c * residue fg z) (circlepath z e)" 
+  have "(fg has_contour_integral c * residue fg z) (circlepath z e)"
     unfolding fg_def using f_holo g_holo
     apply (intro base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
     by (auto intro:holomorphic_intros)
   moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)"
     unfolding fg_def using f_holo g_holo
-    by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])      
+    by (auto intro: has_contour_integral_add base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c_def])
   ultimately have "c*(residue f z + residue g z) = c * residue fg z"
     using has_contour_integral_unique by (auto simp add:distrib_left)
-  thus ?thesis unfolding fg_def 
+  thus ?thesis unfolding fg_def
     by (auto simp add:c_def)
 qed
 
-  
+
 lemma residue_lmul:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
   shows "residue (\<lambda>z. c * (f z)) z= c * residue f z"
 proof (cases "c=0")
   case True
   thus ?thesis using residue_const by auto
-next  
+next
   case False
   def c'\<equiv>"2 * pi * \<i>"
   def f'\<equiv>"(\<lambda>z. c * (f z))"
-  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> 
+  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
     using open_contains_cball_eq by blast
   have "(f' has_contour_integral c' * residue f' z) (circlepath z e)"
     unfolding f'_def using f_holo
@@ -2351,34 +2351,34 @@
     by (auto intro:holomorphic_intros)
   moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)"
     unfolding f'_def using f_holo
-    by (auto intro: has_contour_integral_lmul 
-      base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])    
+    by (auto intro: has_contour_integral_lmul
+      base_residue[OF \<open>open s\<close> \<open>z\<in>s\<close> \<open>e>0\<close> _ e_cball,folded c'_def])
   ultimately have "c' * residue f' z  = c * (c' * residue f z)"
-    using has_contour_integral_unique by auto  
+    using has_contour_integral_unique by auto
   thus ?thesis unfolding f'_def c'_def using False
     by (auto simp add:field_simps)
 qed
 
 lemma residue_rmul:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
   shows "residue (\<lambda>z. (f z) * c) z= residue f z * c"
 using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps)
 
 lemma residue_div:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
   shows "residue (\<lambda>z. (f z) / c) z= residue f z / c "
 using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps)
 
 lemma residue_neg:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
   shows "residue (\<lambda>z. - (f z)) z= - residue f z"
 using residue_lmul[OF assms,of "-1"] by auto
 
 lemma residue_diff:
-  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}" 
+  assumes "open s" "z \<in> s" and f_holo: "f holomorphic_on s - {z}"
       and g_holo:"g holomorphic_on s - {z}"
   shows "residue (\<lambda>z. f z - g z) z= residue f z - residue g z"
-using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)] 
+using residue_add[OF assms(1,2,3),of "\<lambda>z. - g z"] residue_neg[OF assms(1,2,4)]
 by (auto intro:holomorphic_intros g_holo)
 
 lemma residue_simple:
@@ -2387,7 +2387,7 @@
 proof -
   define c where "c \<equiv> 2 * pi * \<i>"
   def f'\<equiv>"\<lambda>w. f w / (w - z)"
-  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close> 
+  obtain e where "e>0" and e_cball:"cball z e \<subseteq> s" using \<open>open s\<close> \<open>z\<in>s\<close>
     using open_contains_cball_eq by blast
   have "(f' has_contour_integral c * f z) (circlepath z e)"
     unfolding f'_def c_def using \<open>e>0\<close> f_holo e_cball
@@ -2407,52 +2407,52 @@
 
 lemma get_integrable_path:
   assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\<in>s-pts" "b\<in>s-pts"
-  obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" 
+  obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b"
     "path_image g \<subseteq> s-pts" "f contour_integrable_on g" using assms
-proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>]) 
+proof (induct arbitrary:s thesis a rule:finite_induct[OF \<open>finite pts\<close>])
   case 1
   obtain g where "valid_path g" "path_image g \<subseteq> s" "pathstart g = a" "pathfinish g = b"
     using connected_open_polynomial_connected[OF \<open>open s\<close>,of a b ] \<open>connected (s - {})\<close>
       valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto
-  moreover have "f contour_integrable_on g" 
+  moreover have "f contour_integrable_on g"
     using contour_integrable_holomorphic_simple[OF _ \<open>open s\<close> \<open>valid_path g\<close> \<open>path_image g \<subseteq> s\<close>,of f]
       \<open>f holomorphic_on s - {}\<close>
     by auto
   ultimately show ?case using "1"(1)[of g] by auto
 next
-  case idt:(2 p pts) 
+  case idt:(2 p pts)
   obtain e where "e>0" and e:"\<forall>w\<in>ball a e. w \<in> s \<and> (w \<noteq> a \<longrightarrow> w \<notin> insert p pts)"
-    using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a]  
+    using finite_ball_avoid[OF \<open>open s\<close> \<open>finite (insert p pts)\<close>, of a]
       \<open>a \<in> s - insert p pts\<close>
     by auto
   define a' where "a' \<equiv> a+e/2"
-  have "a'\<in>s-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>  
+  have "a'\<in>s-{p} -pts"  using e[rule_format,of "a+e/2"] \<open>e>0\<close>
     by (auto simp add:dist_complex_def a'_def)
-  then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" 
-    "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'"    
-    using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) 
+  then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b"
+    "path_image g' \<subseteq> s - {p} - pts" "f contour_integrable_on g'"
+    using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1)
     by (metis Diff_insert2 open_delete)
   define g where "g \<equiv> linepath a a' +++ g'"
   have "valid_path g" unfolding g_def by (auto intro: valid_path_join)
-  moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto 
-  moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def 
+  moreover have "pathstart g = a" and  "pathfinish g = b" unfolding g_def by auto
+  moreover have "path_image g \<subseteq> s - insert p pts" unfolding g_def
     proof (rule subset_path_image_join)
-      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> 
+      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
         by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
       then show "path_image (linepath a a') \<subseteq> s - insert p pts" using e idt(9)
         by auto
     next
       show "path_image g' \<subseteq> s - insert p pts" using g'(4) by blast
     qed
-  moreover have "f contour_integrable_on g"  
+  moreover have "f contour_integrable_on g"
     proof -
-      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close> 
+      have "closed_segment a a' \<subseteq> ball a e" using \<open>e>0\<close>
         by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute)
-      then have "continuous_on (closed_segment a a') f" 
+      then have "continuous_on (closed_segment a a') f"
         using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)]
         apply (elim continuous_on_subset)
         by auto
-      then have "f contour_integrable_on linepath a a'" 
+      then have "f contour_integrable_on linepath a a'"
         using contour_integrable_continuous_linepath by auto
       then show ?thesis unfolding g_def
         apply (rule contour_integrable_joinI)
@@ -2463,16 +2463,16 @@
 
 lemma Cauchy_theorem_aux:
   assumes "open s" "connected (s-pts)" "finite pts" "pts \<subseteq> s" "f holomorphic_on s-pts"
-          "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts" 
-          "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" 
+          "valid_path g" "pathfinish g = pathstart g" "path_image g \<subseteq> s-pts"
+          "\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
           "\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
   shows "contour_integral g f = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
     using assms
-proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>]) 
+proof (induct arbitrary:s g rule:finite_induct[OF \<open>finite pts\<close>])
   case 1
   then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique)
 next
-  case (2 p pts) 
+  case (2 p pts)
   note fin[simp] = \<open>finite (insert p pts)\<close>
     and connected = \<open>connected (s - insert p pts)\<close>
     and valid[simp] = \<open>valid_path g\<close>
@@ -2484,12 +2484,12 @@
   have "h p>0" and "p\<in>s"
     and h_p: "\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> insert p pts)"
     using h \<open>insert p pts \<subseteq> s\<close> by auto
-  obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" 
-      "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg" 
+  obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p"
+      "path_image pg \<subseteq> s-insert p pts" "f contour_integrable_on pg"
     proof -
-      have "p + h p\<in>cball p (h p)" using h[rule_format,of p] 
+      have "p + h p\<in>cball p (h p)" using h[rule_format,of p]
         by (simp add: \<open>p \<in> s\<close> dist_norm)
-      then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> 
+      then have "p + h p \<in> s - insert p pts" using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close>
         by fastforce
       moreover have "pathstart g \<in> s - insert p pts " using path_img by auto
       ultimately show ?thesis
@@ -2497,32 +2497,32 @@
         by blast
     qed
   obtain n::int where "n=winding_number g p"
-    using integer_winding_number[OF _ g_loop,of p] valid path_img 
+    using integer_winding_number[OF _ g_loop,of p] valid path_img
     by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path)
-  define p_circ where "p_circ \<equiv> circlepath p (h p)" 
+  define p_circ where "p_circ \<equiv> circlepath p (h p)"
   define p_circ_pt where "p_circ_pt \<equiv> linepath (p+h p) (p+h p)"
   define n_circ where "n_circ \<equiv> \<lambda>n. (op +++ p_circ ^^ n) p_circ_pt"
   define cp where "cp \<equiv> if n\<ge>0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))"
-  have n_circ:"valid_path (n_circ k)"    
+  have n_circ:"valid_path (n_circ k)"
       "winding_number (n_circ k) p = k"
-      "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" 
+      "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p"
       "path_image (n_circ k) =  (if k=0 then {p + h p} else sphere p (h p))"
       "p \<notin> path_image (n_circ k)"
       "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number (n_circ k) p'=0 \<and> p'\<notin>path_image (n_circ k)"
       "f contour_integrable_on (n_circ k)"
       "contour_integral (n_circ k) f = k *  contour_integral p_circ f"
-      for k 
-    proof (induct k) 
-      case 0 
-      show "valid_path (n_circ 0)" 
+      for k
+    proof (induct k)
+      case 0
+      show "valid_path (n_circ 0)"
         and "path_image (n_circ 0) =  (if 0=0 then {p + h p} else sphere p (h p))"
-        and "winding_number (n_circ 0) p = of_nat 0" 
+        and "winding_number (n_circ 0) p = of_nat 0"
         and "pathstart (n_circ 0) = p + h p"
         and "pathfinish (n_circ 0) = p + h p"
         and "p \<notin> path_image (n_circ 0)"
         unfolding n_circ_def p_circ_pt_def using \<open>h p > 0\<close>
         by (auto simp add: dist_norm)
-      show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p' 
+      show "winding_number (n_circ 0) p'=0 \<and> p'\<notin>path_image (n_circ 0)" when "p'\<notin>s- pts" for p'
         unfolding n_circ_def p_circ_pt_def
         apply (auto intro!:winding_number_trivial)
         by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+
@@ -2534,51 +2534,51 @@
     next
       case (Suc k)
       have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto
-      have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" 
+      have pcirc:"p \<notin> path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)"
         using Suc(3) unfolding p_circ_def using \<open>h p > 0\<close> by (auto simp add: p_circ_def)
-      have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts" 
+      have pcirc_image:"path_image p_circ \<subseteq> s - insert p pts"
         proof -
           have "path_image p_circ \<subseteq> cball p (h p)" using \<open>0 < h p\<close> p_circ_def by auto
           then show ?thesis using h_p pcirc(1) by auto
         qed
       have pcirc_integrable:"f contour_integrable_on p_circ"
-        by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] 
-          contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on 
+        by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def]
+          contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on
           holomorphic_on_subset[OF holo])
-      show "valid_path (n_circ (Suc k))"   
+      show "valid_path (n_circ (Suc k))"
         using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto
-      show "path_image (n_circ (Suc k)) 
-          = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" 
+      show "path_image (n_circ (Suc k))
+          = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))"
         proof -
-          have "path_image p_circ = sphere p (h p)" 
+          have "path_image p_circ = sphere p (h p)"
             unfolding p_circ_def using \<open>0 < h p\<close> by auto
           then show ?thesis unfolding n_Suc  using Suc.hyps(5)  \<open>h p>0\<close>
             by (auto simp add:  path_image_join[OF pcirc(3)]  dist_norm)
         qed
       then show "p \<notin> path_image (n_circ (Suc k))" using \<open>h p>0\<close> by auto
-      show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" 
+      show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)"
         proof -
-          have "winding_number p_circ p = 1" 
+          have "winding_number p_circ p = 1"
             by (simp add: \<open>h p > 0\<close> p_circ_def winding_number_circlepath_centre)
           moreover have "p \<notin> path_image (n_circ k)" using Suc(5) \<open>h p>0\<close> by auto
-          then have "winding_number (p_circ +++ n_circ k) p 
+          then have "winding_number (p_circ +++ n_circ k) p
               = winding_number p_circ p + winding_number (n_circ k) p"
-            using  valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc 
+            using  valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc
             apply (intro winding_number_join)
             by auto
-          ultimately show ?thesis using Suc(2) unfolding n_circ_def 
+          ultimately show ?thesis using Suc(2) unfolding n_circ_def
             by auto
-        qed 
+        qed
       show "pathstart (n_circ (Suc k)) = p + h p"
         by (simp add: n_circ_def p_circ_def)
       show "pathfinish (n_circ (Suc k)) = p + h p"
-        using Suc(4) unfolding n_circ_def by auto      
+        using Suc(4) unfolding n_circ_def by auto
       show "winding_number (n_circ (Suc k)) p'=0 \<and>  p'\<notin>path_image (n_circ (Suc k))" when "p'\<notin>s-pts" for p'
         proof -
           have " p' \<notin> path_image p_circ" using \<open>p \<in> s\<close> h p_circ_def that using pcirc_image by blast
-          moreover have "p' \<notin> path_image (n_circ k)" 
+          moreover have "p' \<notin> path_image (n_circ k)"
             using Suc.hyps(7) that by blast
-          moreover have "winding_number p_circ p' = 0" 
+          moreover have "winding_number p_circ p' = 0"
             proof -
               have "path_image p_circ \<subseteq> cball p (h p)"
                 using h unfolding p_circ_def using \<open>p \<in> s\<close> by fastforce
@@ -2597,28 +2597,28 @@
         by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)])
       show "contour_integral (n_circ (Suc k)) f = (Suc k) *  contour_integral p_circ f"
         unfolding n_Suc
-        by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] 
+        by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]
           Suc(9) algebra_simps)
     qed
   have cp[simp]:"pathstart cp = p + h p"  "pathfinish cp = p + h p"
-         "valid_path cp" "path_image cp \<subseteq> s - insert p pts" 
+         "valid_path cp" "path_image cp \<subseteq> s - insert p pts"
          "winding_number cp p = - n"
          "\<And>p'. p'\<notin>s - pts \<Longrightarrow> winding_number cp p'=0 \<and> p' \<notin> path_image cp"
          "f contour_integrable_on cp"
          "contour_integral cp f = - n * contour_integral p_circ f"
-    proof - 
+    proof -
       show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp"
         using n_circ unfolding cp_def by auto
     next
-      have "sphere p (h p) \<subseteq>  s - insert p pts" 
+      have "sphere p (h p) \<subseteq>  s - insert p pts"
         using h[rule_format,of p] \<open>insert p pts \<subseteq> s\<close> by force
-      moreover  have "p + complex_of_real (h p) \<in> s - insert p pts" 
+      moreover  have "p + complex_of_real (h p) \<in> s - insert p pts"
         using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE)
       ultimately show "path_image cp \<subseteq>  s - insert p pts" unfolding cp_def
         using n_circ(5)  by auto
     next
-      show "winding_number cp p = - n" 
-        unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close> 
+      show "winding_number cp p = - n"
+        unfolding cp_def using winding_number_reversepath n_circ \<open>h p>0\<close>
         by (auto simp: valid_path_imp_path)
     next
       show "winding_number cp p'=0 \<and> p' \<notin> path_image cp" when "p'\<notin>s - pts" for p'
@@ -2636,23 +2636,23 @@
     qed
   def g'\<equiv>"g +++ pg +++ cp +++ (reversepath pg)"
   have "contour_integral g' f = (\<Sum>p\<in>pts. winding_number g' p * contour_integral (circlepath p (h p)) f)"
-    proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ]) 
+    proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \<open>finite pts\<close> ])
       show "connected (s - {p} - pts)" using connected by (metis Diff_insert2)
       show "open (s - {p})" using \<open>open s\<close> by auto
-      show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close>  by blast 
+      show " pts \<subseteq> s - {p}" using \<open>insert p pts \<subseteq> s\<close> \<open> p \<notin> pts\<close>  by blast
       show "f holomorphic_on s - {p} - pts" using holo \<open>p \<notin> pts\<close> by (metis Diff_insert2)
-      show "valid_path g'" 
+      show "valid_path g'"
         unfolding g'_def cp_def using n_circ valid pg g_loop
         by (auto intro!:valid_path_join )
       show "pathfinish g' = pathstart g'"
         unfolding g'_def cp_def using pg(2) by simp
-      show "path_image g' \<subseteq> s - {p} - pts" 
+      show "path_image g' \<subseteq> s - {p} - pts"
         proof -
           def s'\<equiv>"s - {p} - pts"
           have s':"s' = s-insert p pts " unfolding s'_def by auto
-          then show ?thesis using path_img pg(4) cp(4)      
-            unfolding g'_def 
-            apply (fold s'_def s')           
+          then show ?thesis using path_img pg(4) cp(4)
+            unfolding g'_def
+            apply (fold s'_def s')
             apply (intro subset_path_image_join)
             by auto
         qed
@@ -2660,12 +2660,12 @@
       show "\<forall>z. z \<notin> s - {p} \<longrightarrow> winding_number g' z = 0"
         proof clarify
           fix z assume z:"z\<notin>s - {p}"
-          have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z 
+          have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z
               + winding_number (pg +++ cp +++ (reversepath pg)) z"
             proof (rule winding_number_join)
               show "path g" using \<open>valid_path g\<close> by (simp add: valid_path_imp_path)
               show "z \<notin> path_image g" using z path_img by auto
-              show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp 
+              show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp
                 by (simp add: valid_path_imp_path)
             next
               have "path_image (pg +++ cp +++ reversepath pg) \<subseteq> s - insert p pts"
@@ -2674,52 +2674,52 @@
             next
               show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto
             qed
-          also have "... = winding_number g z + (winding_number pg z 
+          also have "... = winding_number g z + (winding_number pg z
               + winding_number (cp +++ (reversepath pg)) z)"
             proof (subst add_left_cancel,rule winding_number_join)
-              show "path pg" and "path (cp +++ reversepath pg)" 
-               and "pathfinish pg = pathstart (cp +++ reversepath pg)" 
+              show "path pg" and "path (cp +++ reversepath pg)"
+               and "pathfinish pg = pathstart (cp +++ reversepath pg)"
                 by (auto simp add: valid_path_imp_path)
               show "z \<notin> path_image pg" using pg(4) z by blast
-              show "z \<notin> path_image (cp +++ reversepath pg)" using z 
-                by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1 
+              show "z \<notin> path_image (cp +++ reversepath pg)" using z
+                by (metis Diff_iff \<open>z \<notin> path_image pg\<close> contra_subsetD cp(4) insertI1
                   not_in_path_image_join path_image_reversepath singletonD)
             qed
-          also have "... = winding_number g z + (winding_number pg z 
-              + (winding_number cp z + winding_number (reversepath pg) z))" 
+          also have "... = winding_number g z + (winding_number pg z
+              + (winding_number cp z + winding_number (reversepath pg) z))"
             apply (auto intro!:winding_number_join simp: valid_path_imp_path)
             apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z)
-            by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)            
+            by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z)
           also have "... = winding_number g z + winding_number cp z"
             apply (subst winding_number_reversepath)
             apply (auto simp: valid_path_imp_path)
             by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z)
-          finally have "winding_number g' z = winding_number g z + winding_number cp z" 
+          finally have "winding_number g' z = winding_number g z + winding_number cp z"
             unfolding g'_def .
-          moreover have "winding_number g z + winding_number cp z = 0" 
+          moreover have "winding_number g z + winding_number cp z = 0"
             using winding z \<open>n=winding_number g p\<close> by auto
           ultimately show "winding_number g' z = 0" unfolding g'_def by auto
         qed
       show "\<forall>pa\<in>s - {p}. 0 < h pa \<and> (\<forall>w\<in>cball pa (h pa). w \<in> s - {p} \<and> (w \<noteq> pa \<longrightarrow> w \<notin> pts))"
         using h by fastforce
     qed
-  moreover have "contour_integral g' f = contour_integral g f 
-      - winding_number g p * contour_integral p_circ f" 
+  moreover have "contour_integral g' f = contour_integral g f
+      - winding_number g p * contour_integral p_circ f"
     proof -
-      have "contour_integral g' f =  contour_integral g f 
+      have "contour_integral g' f =  contour_integral g f
         + contour_integral (pg +++ cp +++ reversepath pg) f"
         unfolding g'_def
-        apply (subst Cauchy_Integral_Thm.contour_integral_join) 
+        apply (subst contour_integral_join)
         by (auto simp add:open_Diff[OF \<open>open s\<close>,OF finite_imp_closed[OF fin]]
-          intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]  
+          intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img]
           contour_integrable_reversepath)
-      also have "... = contour_integral g f + contour_integral pg f 
+      also have "... = contour_integral g f + contour_integral pg f
           + contour_integral (cp +++ reversepath pg) f"
-        apply (subst Cauchy_Integral_Thm.contour_integral_join)
+        apply (subst contour_integral_join)
         by (auto simp add:contour_integrable_reversepath)
-      also have "... = contour_integral g f + contour_integral pg f 
+      also have "... = contour_integral g f + contour_integral pg f
           + contour_integral cp f + contour_integral (reversepath pg) f"
-        apply (subst Cauchy_Integral_Thm.contour_integral_join)
+        apply (subst contour_integral_join)
         by (auto simp add:contour_integrable_reversepath)
       also have "... = contour_integral g f + contour_integral cp f"
         using contour_integral_reversepath
@@ -2728,21 +2728,21 @@
         using \<open>n=winding_number g p\<close> by auto
       finally show ?thesis .
     qed
-  moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p' 
+  moreover have "winding_number g' p' = winding_number g p'" when "p'\<in>pts" for p'
     proof -
-      have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp" 
-        using "2.prems"(8) that 
+      have [simp]: "p' \<notin> path_image g" "p' \<notin> path_image pg" "p'\<notin>path_image cp"
+        using "2.prems"(8) that
         apply blast
         apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that)
         by (meson DiffD2 cp(4) set_rev_mp subset_insertI that)
-      have "winding_number g' p' = winding_number g p' 
-          + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def 
+      have "winding_number g' p' = winding_number g p'
+          + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def
         apply (subst winding_number_join)
         apply (simp_all add: valid_path_imp_path)
         apply (intro not_in_path_image_join)
         by auto
       also have "... = winding_number g p' + winding_number pg p'
-          + winding_number (cp +++ reversepath pg) p'" 
+          + winding_number (cp +++ reversepath pg) p'"
         apply (subst winding_number_join)
         apply (simp_all add: valid_path_imp_path)
         apply (intro not_in_path_image_join)
@@ -2757,16 +2757,16 @@
       also have "... = winding_number g p'" using that by auto
       finally show ?thesis .
     qed
-  ultimately show ?case unfolding p_circ_def 
-    apply (subst (asm) setsum.cong[OF refl, 
+  ultimately show ?case unfolding p_circ_def
+    apply (subst (asm) setsum.cong[OF refl,
         of pts _ "\<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"])
-    by (auto simp add:setsum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)         
-qed  
+    by (auto simp add:setsum.insert[OF \<open>finite pts\<close> \<open>p\<notin>pts\<close>] algebra_simps)
+qed
 
 lemma Cauchy_theorem_singularities:
-  assumes "open s" "connected s" "finite pts" and 
+  assumes "open s" "connected s" "finite pts" and
           holo:"f holomorphic_on s-pts" and
-          "valid_path g" and 
+          "valid_path g" and
           loop:"pathfinish g = pathstart g" and
           "path_image g \<subseteq> s-pts" and
           homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" and
@@ -2775,32 +2775,32 @@
     (is "?L=?R")
 proof -
   define circ where "circ \<equiv> \<lambda>p. winding_number g p * contour_integral (circlepath p (h p)) f"
-  define pts1 where "pts1 \<equiv> pts \<inter> s" 
+  define pts1 where "pts1 \<equiv> pts \<inter> s"
   define pts2 where "pts2 \<equiv> pts - pts1"
-  have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s" 
+  have "pts=pts1 \<union> pts2" "pts1 \<inter> pts2 = {}" "pts2 \<inter> s={}" "pts1\<subseteq>s"
     unfolding pts1_def pts2_def by auto
   have "contour_integral g f =  (\<Sum>p\<in>pts1. circ p)" unfolding circ_def
     proof (rule Cauchy_theorem_aux[OF \<open>open s\<close> _ _ \<open>pts1\<subseteq>s\<close> _ \<open>valid_path g\<close> loop _ homo])
       have "finite pts1" unfolding pts1_def using \<open>finite pts\<close> by auto
-      then show "connected (s - pts1)" 
+      then show "connected (s - pts1)"
         using \<open>open s\<close> \<open>connected s\<close> connected_open_delete_finite[of s] by auto
     next
       show "finite pts1" using \<open>pts = pts1 \<union> pts2\<close> assms(3) by auto
       show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def)
-      show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto 
+      show "path_image g \<subseteq> s - pts1" using assms(7) pts1_def by auto
       show "\<forall>p\<in>s. 0 < h p \<and> (\<forall>w\<in>cball p (h p). w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> pts1))"
         by (simp add: avoid pts1_def)
     qed
   moreover have "setsum circ pts2=0"
     proof -
-      have "winding_number g p=0" when "p\<in>pts2" for p 
+      have "winding_number g p=0" when "p\<in>pts2" for p
         using  \<open>pts2 \<inter> s={}\<close> that homo[rule_format,of p] by auto
       thus ?thesis unfolding circ_def
         apply (intro setsum.neutral)
         by auto
     qed
   moreover have "?R=setsum circ pts1 + setsum circ pts2"
-    unfolding circ_def 
+    unfolding circ_def
     using setsum.union_disjoint[OF _ _ \<open>pts1 \<inter> pts2 = {}\<close>] \<open>finite pts\<close> \<open>pts=pts1 \<union> pts2\<close>
     by blast
   ultimately show ?thesis
@@ -2811,34 +2811,34 @@
 lemma Residue_theorem:
   fixes s pts::"complex set" and f::"complex \<Rightarrow> complex"
     and g::"real \<Rightarrow> complex"
-  assumes "open s" "connected s" "finite pts" and 
+  assumes "open s" "connected s" "finite pts" and
           holo:"f holomorphic_on s-pts" and
-          "valid_path g" and 
+          "valid_path g" and
           loop:"pathfinish g = pathstart g" and
           "path_image g \<subseteq> s-pts" and
-          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0" 
+          homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z  = 0"
   shows "contour_integral g f = 2 * pi * \<i> *(\<Sum>p\<in>pts. winding_number g p * residue f p)"
 proof -
   define c where "c \<equiv>  2 * pi * \<i>"
-  obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))" 
+  obtain h where avoid:"\<forall>p\<in>s. h p>0 \<and> (\<forall>w\<in>cball p (h p). w\<in>s \<and> (w\<noteq>p \<longrightarrow> w \<notin> pts))"
     using finite_cball_avoid[OF \<open>open s\<close> \<open>finite pts\<close>] by metis
-  have "contour_integral g f 
+  have "contour_integral g f
       = (\<Sum>p\<in>pts. winding_number g p * contour_integral (circlepath p (h p)) f)"
     using Cauchy_theorem_singularities[OF assms avoid] .
   also have "... = (\<Sum>p\<in>pts.  c * winding_number g p * residue f p)"
     proof (intro setsum.cong)
       show "pts = pts" by simp
     next
-      fix x assume "x \<in> pts" 
-      show "winding_number g x * contour_integral (circlepath x (h x)) f 
-          = c * winding_number g x * residue f x"   
+      fix x assume "x \<in> pts"
+      show "winding_number g x * contour_integral (circlepath x (h x)) f
+          = c * winding_number g x * residue f x"
         proof (cases "x\<in>s")
           case False
           then have "winding_number g x=0" using homo by auto
           thus ?thesis by auto
         next
           case True
-          have "contour_integral (circlepath x (h x)) f = c* residue f x"  
+          have "contour_integral (circlepath x (h x)) f = c* residue f x"
             using \<open>x\<in>pts\<close> \<open>finite pts\<close> avoid[rule_format,OF True]
             apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def])
             by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \<open>open s\<close> finite_imp_closed])
@@ -2846,9 +2846,9 @@
         qed
     qed
   also have "... = c * (\<Sum>p\<in>pts. winding_number g p * residue f p)"
-    by (simp add: setsum_right_distrib algebra_simps) 
+    by (simp add: setsum_right_distrib algebra_simps)
   finally show ?thesis unfolding c_def .
-qed    
+qed
 
 subsection \<open>The argument principle\<close>
 
@@ -2858,13 +2858,13 @@
 lemma is_pole_tendsto:
   fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)"
   shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)"
-unfolding is_pole_def 
+unfolding is_pole_def
 by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)
 
 lemma is_pole_inverse_holomorphic:
   assumes "open s"
-    and f_holo:"f holomorphic_on (s-{z})" 
-    and pole:"is_pole f z" 
+    and f_holo:"f holomorphic_on (s-{z})"
+    and pole:"is_pole f z"
     and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0"
   shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s"
 proof -
@@ -2875,7 +2875,7 @@
   moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto
   hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
     by (auto elim!:continuous_on_inverse simp add:non_z)
-  hence "continuous_on (s-{z}) g" unfolding g_def    
+  hence "continuous_on (s-{z}) g" unfolding g_def
     apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"])
     by auto
   ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close>
@@ -2887,17 +2887,17 @@
     apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"])
     by (auto simp add:g_def)
   ultimately show ?thesis unfolding g_def using \<open>open s\<close>
-    by (auto elim!: no_isolated_singularity)    
+    by (auto elim!: no_isolated_singularity)
 qed
 
 
 (*order of the zero of f at z*)
 definition zorder::"(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> nat" where
-  "zorder f z = (THE n. n>0 \<and> (\<exists>h r. r>0 \<and> h holomorphic_on cball z r 
+  "zorder f z = (THE n. n>0 \<and> (\<exists>h r. r>0 \<and> h holomorphic_on cball z r
                     \<and> (\<forall>w\<in>cball z r. f w =  h w * (w-z)^n \<and> h w \<noteq>0)))"
 
 definition zer_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where
-  "zer_poly f z = (SOME h. \<exists>r . r>0 \<and> h holomorphic_on cball z r 
+  "zer_poly f z = (SOME h. \<exists>r . r>0 \<and> h holomorphic_on cball z r
                     \<and> (\<forall>w\<in>cball z r. f w =  h w * (w-z)^(zorder f z) \<and> h w \<noteq>0))"
 
 (*order of the pole of f at z*)
@@ -2905,25 +2905,25 @@
   "porder f z = (let f'=(\<lambda>x. if x=z then 0 else inverse (f x)) in zorder f' z)"
 
 definition pol_poly::"[complex \<Rightarrow> complex,complex]\<Rightarrow>complex \<Rightarrow> complex" where
-  "pol_poly f z = (let f'=(\<lambda> x. if x=z then 0 else inverse (f x)) 
+  "pol_poly f z = (let f'=(\<lambda> x. if x=z then 0 else inverse (f x))
       in inverse o zer_poly f' z)"
 
 
-lemma holomorphic_factor_zero_unique: 
-  fixes f::"complex \<Rightarrow> complex" and z::complex and r::real    
+lemma holomorphic_factor_zero_unique:
+  fixes f::"complex \<Rightarrow> complex" and z::complex and r::real
   assumes "r>0"
     and asm:"\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0"
     and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r"
-  shows "n=m" 
+  shows "n=m"
 proof -
   have "n>m \<Longrightarrow> False"
     proof -
       assume "n>m"
-      have "(h \<longlongrightarrow> 0) (at z within ball z r)" 
-        proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (n - m) * g w"]) 
+      have "(h \<longlongrightarrow> 0) (at z within ball z r)"
+        proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (n - m) * g w"])
           have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> h w = (w-z)^(n-m) * g w" using \<open>n>m\<close> asm
             by (auto simp add:field_simps power_diff)
-          then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> 
+          then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
             \<Longrightarrow> (x' - z) ^ (n - m) * g x' = h x'" for x' by auto
         next
           define F where "F \<equiv> at z within ball z r"
@@ -2932,16 +2932,16 @@
           moreover have "continuous F f'" unfolding f'_def F_def
             by (intro continuous_intros)
           ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
-            by (simp add: continuous_within)            
+            by (simp add: continuous_within)
           moreover have "(g \<longlongrightarrow> g z) F"
             using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close>
             unfolding F_def by auto
           ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
         qed
-      moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)" 
-        using holomorphic_on_imp_continuous_on[OF h_holo] 
-        by (auto simp add:continuous_on_def \<open>r>0\<close>) 
-      moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> 
+      moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)"
+        using holomorphic_on_imp_continuous_on[OF h_holo]
+        by (auto simp add:continuous_on_def \<open>r>0\<close>)
+      moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close>
         by (auto simp add:trivial_limit_within islimpt_ball)
       ultimately have "h z=0" by (auto intro: tendsto_unique)
       thus False using asm \<open>r>0\<close> by auto
@@ -2949,11 +2949,11 @@
   moreover have "m>n \<Longrightarrow> False"
     proof -
       assume "m>n"
-      have "(g \<longlongrightarrow> 0) (at z within ball z r)" 
-        proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (m - n) * h w"]) 
+      have "(g \<longlongrightarrow> 0) (at z within ball z r)"
+        proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) ^ (m - n) * h w"])
           have "\<forall>w\<in>ball z r. w\<noteq>z \<longrightarrow> g w = (w-z)^(m-n) * h w" using \<open>m>n\<close> asm
             by (auto simp add:field_simps power_diff)
-          then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk> 
+          then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
             \<Longrightarrow> (x' - z) ^ (m - n) * h x' = g x'" for x' by auto
         next
           define F where "F \<equiv> at z within ball z r"
@@ -2962,16 +2962,16 @@
           moreover have "continuous F f'" unfolding f'_def F_def
             by (intro continuous_intros)
           ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
-            by (simp add: continuous_within)            
+            by (simp add: continuous_within)
           moreover have "(h \<longlongrightarrow> h z) F"
             using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close>
             unfolding F_def by auto
           ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
         qed
-      moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)" 
-        using holomorphic_on_imp_continuous_on[OF g_holo] 
-        by (auto simp add:continuous_on_def \<open>r>0\<close>) 
-      moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close> 
+      moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)"
+        using holomorphic_on_imp_continuous_on[OF g_holo]
+        by (auto simp add:continuous_on_def \<open>r>0\<close>)
+      moreover have "at z within ball z r \<noteq> bot" using \<open>r>0\<close>
         by (auto simp add:trivial_limit_within islimpt_ball)
       ultimately have "g z=0" by (auto intro: tendsto_unique)
       thus False using asm \<open>r>0\<close> by auto
@@ -2982,62 +2982,62 @@
 
 lemma holomorphic_factor_zero_Ex1:
   assumes "open s" "connected s" "z \<in> s" and
-        holo:"f holomorphic_on s" 
+        holo:"f holomorphic_on s"
         and "f z = 0" and "\<exists>w\<in>s. f w \<noteq> 0"
   shows "\<exists>!n. \<exists>g r. 0 < n \<and> 0 < r \<and>
-                g holomorphic_on cball z r 
+                g holomorphic_on cball z r
                 \<and> (\<forall>w\<in>cball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0)"
 proof (rule ex_ex1I)
-  obtain g r n where "0 < n" "0 < r" "ball z r \<subseteq> s" and 
+  obtain g r n where "0 < n" "0 < r" "ball z r \<subseteq> s" and
           g:"g holomorphic_on ball z r"
-          "\<And>w. w \<in> ball z r \<Longrightarrow> f w = (w - z) ^ n * g w" 
+          "\<And>w. w \<in> ball z r \<Longrightarrow> f w = (w - z) ^ n * g w"
           "\<And>w. w \<in> ball z r \<Longrightarrow> g w \<noteq> 0"
     using holomorphic_factor_zero_nonconstant[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> \<open>f z=0\<close>]
     by (metis assms(3) assms(5) assms(6))
   def r'\<equiv>"r/2"
   have "cball z r' \<subseteq> ball z r" unfolding r'_def by (simp add: \<open>0 < r\<close> cball_subset_ball_iff)
-  hence "cball z r' \<subseteq> s" "g holomorphic_on cball z r'" 
+  hence "cball z r' \<subseteq> s" "g holomorphic_on cball z r'"
       "(\<forall>w\<in>cball z r'. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)"
     using g \<open>ball z r \<subseteq> s\<close> by auto
   moreover have "r'>0" unfolding r'_def using \<open>0<r\<close> by auto
-  ultimately show "\<exists>n g r. 0 < n \<and> 0 < r  \<and> g holomorphic_on cball z r 
+  ultimately show "\<exists>n g r. 0 < n \<and> 0 < r  \<and> g holomorphic_on cball z r
           \<and> (\<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0)"
     apply (intro exI[of _ n] exI[of _ g] exI[of _ r'])
     by (simp add:\<open>0 < n\<close>)
 next
-  fix m n 
-  define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0"  
+  fix m n
+  define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r. f w = (w - z) ^ n * g w \<and> g w \<noteq> 0"
   assume n_asm:"\<exists>g r1. 0 < n \<and> 0 < r1 \<and> g holomorphic_on cball z r1 \<and> fac n g r1"
      and m_asm:"\<exists>h r2. 0 < m \<and> 0 < r2  \<and> h holomorphic_on cball z r2 \<and> fac m h r2"
-  obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1" 
+  obtain g r1 where "0 < n" "0 < r1" and g_holo: "g holomorphic_on cball z r1"
     and "fac n g r1" using n_asm by auto
-  obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2"  
+  obtain h r2 where "0 < m" "0 < r2" and h_holo: "h holomorphic_on cball z r2"
     and "fac m h r2" using m_asm by auto
   define r where "r \<equiv> min r1 r2"
   have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
-  moreover have "\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0" 
+  moreover have "\<forall>w\<in>ball z r. f w = (w-z)^n * g w \<and> g w\<noteq>0 \<and> f w = (w - z)^m * h w \<and> h w\<noteq>0"
     using \<open>fac m h r2\<close> \<open>fac n g r1\<close>   unfolding fac_def r_def
     by fastforce
-  ultimately show "m=n" using g_holo h_holo 
+  ultimately show "m=n" using g_holo h_holo
     apply (elim holomorphic_factor_zero_unique[of r z f n g m h,symmetric,rotated])
     by (auto simp add:r_def)
 qed
 
-lemma zorder_exist:  
+lemma zorder_exist:
   fixes f::"complex \<Rightarrow> complex" and z::complex
   defines "n\<equiv>zorder f z" and "h\<equiv>zer_poly f z"
-  assumes  "open s" "connected s" "z\<in>s" 
+  assumes  "open s" "connected s" "z\<in>s"
     and holo: "f holomorphic_on s"
     and  "f z=0" "\<exists>w\<in>s. f w\<noteq>0"
-  shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r 
+  shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r
     \<and> (\<forall>w\<in>cball z r. f w  = h w * (w-z)^n \<and> h w \<noteq>0) "
 proof -
-  define P where "P \<equiv> \<lambda>h r n. r>0 \<and> h holomorphic_on cball z r 
+  define P where "P \<equiv> \<lambda>h r n. r>0 \<and> h holomorphic_on cball z r
     \<and> (\<forall>w\<in>cball z r. ( f w  = h w * (w-z)^n) \<and> h w \<noteq>0)"
   have "(\<exists>!n. n>0 \<and> (\<exists> h r. P h r n))"
     proof -
       have "\<exists>!n. \<exists>h r. n>0 \<and> P h r n"
-        using holomorphic_factor_zero_Ex1[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo \<open>f z=0\<close> 
+        using holomorphic_factor_zero_Ex1[OF \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close> holo \<open>f z=0\<close>
           \<open>\<exists>w\<in>s. f w\<noteq>0\<close>] unfolding P_def
         apply (subst mult.commute)
         by auto
@@ -3045,77 +3045,77 @@
     qed
   moreover have n:"n=(THE n. n>0 \<and> (\<exists>h r. P h r n))"
     unfolding n_def zorder_def P_def by simp
-  ultimately have "n>0 \<and> (\<exists>h r. P h r n)"  
+  ultimately have "n>0 \<and> (\<exists>h r. P h r n)"
     apply (drule_tac theI')
     by simp
   then have "n>0" and "\<exists>h r. P h r n" by auto
   moreover have "h=(SOME h. \<exists>r. P h r n)"
     unfolding h_def P_def zer_poly_def[of f z,folded n_def P_def] by simp
-  ultimately have "\<exists>r. P h r n" 
+  ultimately have "\<exists>r. P h r n"
     apply (drule_tac someI_ex)
     by simp
   then obtain r1 where "P h r1 n" by auto
-  obtain r2 where "r2>0" "cball z r2 \<subseteq> s" 
+  obtain r2 where "r2>0" "cball z r2 \<subseteq> s"
     using assms(3) assms(5) open_contains_cball_eq by blast
   define r3 where "r3 \<equiv> min r1 r2"
   have "P h r3 n" using \<open>P h r1 n\<close> \<open>r2>0\<close> unfolding P_def r3_def
     by auto
   moreover have "cball z r3 \<subseteq> s" using \<open>cball z r2 \<subseteq> s\<close> unfolding r3_def by auto
   ultimately show ?thesis using \<open>n>0\<close> unfolding P_def by auto
-qed    
+qed
 
-lemma porder_exist:  
+lemma porder_exist:
   fixes f::"complex \<Rightarrow> complex" and z::complex
   defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z"
-  assumes "open s" "z \<in> s" 
+  assumes "open s" "z \<in> s"
     and holo:"f holomorphic_on s-{z}"
     and "is_pole f z"
-  shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r 
+  shows "\<exists>r. n>0 \<and> r>0 \<and> cball z r \<subseteq> s \<and> h holomorphic_on cball z r
     \<and> (\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w  = h w / (w-z)^n) \<and> h w \<noteq>0)"
 proof -
   obtain e where "e>0" and e_ball:"ball z e \<subseteq> s"and e_def: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0"
-    proof - 
-      have "\<forall>\<^sub>F z in at z. f z \<noteq> 0" 
-        using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def  
+    proof -
+      have "\<forall>\<^sub>F z in at z. f z \<noteq> 0"
+        using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def
         by auto
       then obtain e1 where "e1>0" and e1_def: "\<forall>x. x \<noteq> z \<and> dist x z < e1 \<longrightarrow> f x \<noteq> 0"
         using eventually_at[of "\<lambda>x. f x\<noteq>0" z,simplified] by auto
       obtain e2 where "e2>0" and "ball z e2 \<subseteq>s" using \<open>open s\<close> \<open>z\<in>s\<close> openE by auto
       define e where "e \<equiv> min e1 e2"
       have "e>0" using \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def by auto
-      moreover have "ball z e \<subseteq> s" unfolding e_def using \<open>ball z e2 \<subseteq> s\<close> by auto 
-      moreover have "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" using e1_def \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def 
+      moreover have "ball z e \<subseteq> s" unfolding e_def using \<open>ball z e2 \<subseteq> s\<close> by auto
+      moreover have "\<forall>x\<in>ball z e-{z}. f x\<noteq>0" using e1_def \<open>e1>0\<close> \<open>e2>0\<close> unfolding e_def
         by (simp add: DiffD1 DiffD2 dist_commute singletonI)
       ultimately show ?thesis using that by auto
     qed
   define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
-  define zo where "zo \<equiv> zorder g z" 
-  define zp where "zp \<equiv> zer_poly g z" 
-  have "\<exists>w\<in>ball z e. g w \<noteq> 0" 
+  define zo where "zo \<equiv> zorder g z"
+  define zp where "zp \<equiv> zer_poly g z"
+  have "\<exists>w\<in>ball z e. g w \<noteq> 0"
     proof -
       obtain w where w:"w\<in>ball z e-{z}" using \<open>0 < e\<close>
-        by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single 
+        by (metis open_ball all_not_in_conv centre_in_ball insert_Diff_single
           insert_absorb not_open_singleton)
       hence "w\<noteq>z" "f w\<noteq>0" using e_def[rule_format,of w] mem_ball
         by (auto simp add:dist_commute)
       then show ?thesis unfolding g_def using w by auto
     qed
-  moreover have "g holomorphic_on ball z e" 
+  moreover have "g holomorphic_on ball z e"
     apply (intro is_pole_inverse_holomorphic[of "ball z e",OF _ _ \<open>is_pole f z\<close> e_def,folded g_def])
     using holo e_ball by auto
   moreover have "g z=0" unfolding g_def by auto
-  ultimately obtain r where "0 < zo" "0 < r" "cball z r \<subseteq> ball z e" 
+  ultimately obtain r where "0 < zo" "0 < r" "cball z r \<subseteq> ball z e"
       and zp_holo: "zp holomorphic_on cball z r" and
       zp_fac: "\<forall>w\<in>cball z r. g w = zp w * (w - z) ^ zo \<and> zp w \<noteq> 0"
-    using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \<open>e>0\<close> 
+    using zorder_exist[of "ball z e" z g,simplified,folded zo_def zp_def] \<open>e>0\<close>
     by auto
-  have n:"n=zo" and h:"h=inverse o zp" 
+  have n:"n=zo" and h:"h=inverse o zp"
     unfolding n_def zo_def porder_def h_def zp_def pol_poly_def g_def by simp_all
   have "h holomorphic_on cball z r"
     using zp_holo zp_fac holomorphic_on_inverse  unfolding h comp_def by blast
   moreover have "\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w  = h w / (w-z)^n) \<and> h w \<noteq>0"
     using zp_fac unfolding h n comp_def g_def
-    by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq 
+    by (metis divide_inverse_commute field_class.field_inverse_zero inverse_inverse_eq
       inverse_mult_distrib mult.commute)
   moreover have "0 < n" unfolding n using \<open>zo>0\<close> by simp
   ultimately show ?thesis using \<open>0 < r\<close> \<open>cball z r \<subseteq> ball z e\<close> e_ball by auto
@@ -3124,13 +3124,13 @@
 lemma residue_porder:
   fixes f::"complex \<Rightarrow> complex" and z::complex
   defines "n \<equiv> porder f z" and "h \<equiv> pol_poly f z"
-  assumes "open s" "z \<in> s" 
+  assumes "open s" "z \<in> s"
     and holo:"f holomorphic_on s - {z}"
     and pole:"is_pole f z"
   shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))"
 proof -
   define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
-  obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> s" and h_holo: "h holomorphic_on cball z r" 
+  obtain r where "0 < n" "0 < r" and r_cball:"cball z r \<subseteq> s" and h_holo: "h holomorphic_on cball z r"
       and h_divide:"(\<forall>w\<in>cball z r. (w\<noteq>z \<longrightarrow> f w = h w / (w - z) ^ n) \<and> h w \<noteq> 0)"
     using porder_exist[OF \<open>open s\<close> \<open>z \<in> s\<close> holo pole, folded n_def h_def] by blast
   have r_nonzero:"\<And>w. w \<in> ball z r - {z} \<Longrightarrow> f w \<noteq> 0"
@@ -3138,7 +3138,7 @@
   define c where "c \<equiv> 2 * pi * \<i>"
   define der_f where "der_f \<equiv> ((deriv ^^ (n - 1)) h z / fact (n-1))"
   def h'\<equiv>"\<lambda>u. h u / (u - z) ^ n"
-  have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" 
+  have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)"
     unfolding h'_def
     proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1",
         folded c_def Suc_pred'[OF \<open>n>0\<close>]])
@@ -3147,9 +3147,9 @@
       show " z \<in> ball z r" using \<open>r>0\<close> by auto
     qed
   then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto
-  then have "(f has_contour_integral c * der_f) (circlepath z r)" 
+  then have "(f has_contour_integral c * der_f) (circlepath z r)"
     proof (elim has_contour_integral_eq)
-      fix x assume "x \<in> path_image (circlepath z r)" 
+      fix x assume "x \<in> path_image (circlepath z r)"
       hence "x\<in>cball z r - {z}" using \<open>r>0\<close> by auto
       then show "h' x = f x" using h_divide unfolding h'_def by auto
     qed
@@ -3163,18 +3163,18 @@
 theorem argument_principle:
   fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set"
   defines "zeros\<equiv>{p. f p=0} - poles"
-  assumes "open s" and 
-          "connected s" and 
-          f_holo:"f holomorphic_on s-poles" and 
+  assumes "open s" and
+          "connected s" and
+          f_holo:"f holomorphic_on s-poles" and
           h_holo:"h holomorphic_on s" and
-          "valid_path g" and 
+          "valid_path g" and
           loop:"pathfinish g = pathstart g" and
           path_img:"path_image g \<subseteq> s - (zeros \<union> poles)" and
           homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and
           finite:"finite (zeros \<union> poles)" and
           poles:"\<forall>p\<in>poles. is_pole f p"
   shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> *
-          ((\<Sum>p\<in>zeros. winding_number g p * h p * zorder f p) 
+          ((\<Sum>p\<in>zeros. winding_number g p * h p * zorder f p)
            - (\<Sum>p\<in>poles. winding_number g p * h p * porder f p))"
     (is "?L=?R")
 proof -
@@ -3183,40 +3183,40 @@
   define cont_pole where "cont_pole \<equiv> \<lambda>ff p e. (ff has_contour_integral - c  * porder f p * h p) (circlepath p e)"
   define cont_zero where "cont_zero \<equiv> \<lambda>ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)"
   define avoid where "avoid \<equiv> \<lambda>p e. \<forall>w\<in>cball p e. w \<in> s \<and> (w \<noteq> p \<longrightarrow> w \<notin> zeros \<union> poles)"
-  have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> cont_pole ff p e) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p e)" 
-      when "p\<in>s" for p 
+  have "\<exists>e>0. avoid p e \<and> (p\<in>poles \<longrightarrow> cont_pole ff p e) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p e)"
+      when "p\<in>s" for p
     proof -
       obtain e1 where "e1>0" and e1_avoid:"avoid p e1"
         using finite_cball_avoid[OF \<open>open s\<close> finite] \<open>p\<in>s\<close> unfolding avoid_def by auto
       have "\<exists>e2>0. cball p e2 \<subseteq> ball p e1 \<and> cont_pole ff p e2"
-        when "p\<in>poles" 
-        proof - 
+        when "p\<in>poles"
+        proof -
           define po where "po \<equiv> porder f p"
           define pp where "pp \<equiv> pol_poly f p"
           def f'\<equiv>"\<lambda>w. pp w / (w - p) ^ po"
-          def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)"         
-          have "f holomorphic_on ball p e1 - {p}" 
+          def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)"
+          have "f holomorphic_on ball p e1 - {p}"
             apply (intro holomorphic_on_subset[OF f_holo])
             using e1_avoid \<open>p\<in>poles\<close> unfolding avoid_def by auto
-          then obtain r where 
+          then obtain r where
               "0 < po" "r>0"
               "cball p r \<subseteq> ball p e1" and
               pp_holo:"pp holomorphic_on cball p r" and
               pp_po:"(\<forall>w\<in>cball p r. (w\<noteq>p \<longrightarrow> f w = pp w / (w - p) ^ po) \<and> pp w \<noteq> 0)"
             using porder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] poles \<open>p\<in>poles\<close>
-            unfolding po_def pp_def 
-            by auto    
+            unfolding po_def pp_def
+            by auto
           define e2 where "e2 \<equiv> r/2"
           have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
-          define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w" 
+          define anal where "anal \<equiv> \<lambda>w. deriv pp w * h w / pp w"
           define prin where "prin \<equiv> \<lambda>w. - of_nat po * h w / (w - p)"
           have "((\<lambda>w.  prin w + anal w) has_contour_integral - c * po * h p) (circlepath p e2)"
             proof (rule  has_contour_integral_add[of _ _ _ _ 0,simplified])
-              have "ball p r \<subseteq> s" 
+              have "ball p r \<subseteq> s"
                 using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast
               then have "cball p e2 \<subseteq> s"
-                using \<open>r>0\<close> unfolding e2_def by auto  
-              then have "(\<lambda>w. - of_nat po * h w) holomorphic_on cball p e2" 
+                using \<open>r>0\<close> unfolding e2_def by auto
+              then have "(\<lambda>w. - of_nat po * h w) holomorphic_on cball p e2"
                 using h_holo
                 by (auto intro!: holomorphic_intros)
               then show "(prin has_contour_integral - c * of_nat po * h p ) (circlepath p e2)"
@@ -3224,7 +3224,7 @@
                   \<open>e2>0\<close>
                 unfolding prin_def
                 by (auto simp add: mult.assoc)
-              have "anal holomorphic_on ball p r" unfolding anal_def 
+              have "anal holomorphic_on ball p r" unfolding anal_def
                 using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close>
                 by (auto intro!: holomorphic_intros)
               then show "(anal has_contour_integral 0) (circlepath p e2)"
@@ -3233,19 +3233,19 @@
             qed
           then have "cont_pole ff' p e2" unfolding cont_pole_def po_def
             proof (elim has_contour_integral_eq)
-              fix w assume "w \<in> path_image (circlepath p e2)" 
+              fix w assume "w \<in> path_image (circlepath p e2)"
               then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
               define wp where "wp \<equiv> w-p"
-              have "wp\<noteq>0" and "pp w \<noteq>0" 
-                unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto 
-              moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po" 
+              have "wp\<noteq>0" and "pp w \<noteq>0"
+                unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto
+              moreover have der_f':"deriv f' w = - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po"
                 proof (rule DERIV_imp_deriv)
                   define der where "der \<equiv> - po * pp w / (w-p)^(po+1) + deriv pp w / (w-p)^po"
                   have po:"po = Suc (po - Suc 0) " using \<open>po>0\<close> by auto
-                  have "(pp has_field_derivative (deriv pp w)) (at w)" 
-                    using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close> 
+                  have "(pp has_field_derivative (deriv pp w)) (at w)"
+                    using DERIV_deriv_iff_has_field_derivative pp_holo \<open>w\<noteq>p\<close>
                       by (meson open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset)
-                  then show "(f' has_field_derivative  der) (at w)" 
+                  then show "(f' has_field_derivative  der) (at w)"
                     using \<open>w\<noteq>p\<close> \<open>po>0\<close> unfolding der_def f'_def
                     apply (auto intro!: derivative_eq_intros simp add:field_simps)
                     apply (subst (4) po)
@@ -3259,16 +3259,16 @@
                 apply (fold wp_def)
                 by (auto simp add:field_simps)
             qed
-          then have "cont_pole ff p e2" unfolding cont_pole_def   
+          then have "cont_pole ff p e2" unfolding cont_pole_def
             proof (elim has_contour_integral_eq)
               fix w assume "w \<in> path_image (circlepath p e2)"
               then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
-              have "deriv f' w =  deriv f w" 
+              have "deriv f' w =  deriv f w"
                 proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
                   show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
                     by (auto intro!: holomorphic_intros)
-                next                
-                  have "ball p e1 - {p} \<subseteq> s - poles" 
+                next
+                  have "ball p e1 - {p} \<subseteq> s - poles"
                     using avoid_def ball_subset_cball e1_avoid
                     by auto
                   then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close>
@@ -3281,16 +3281,16 @@
                   show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto
                 next
                   fix x assume "x \<in> ball p r - {p}"
-                  then show "f' x = f x" 
+                  then show "f' x = f x"
                     using pp_po unfolding f'_def by auto
                 qed
-              moreover have " f' w  =  f w " 
-                using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close> 
+              moreover have " f' w  =  f w "
+                using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po \<open>w\<noteq>p\<close>
                 unfolding f'_def by auto
               ultimately show "ff' w = ff w"
                 unfolding ff'_def ff_def by simp
             qed
-          moreover have "cball p e2 \<subseteq> ball p e1" 
+          moreover have "cball p e2 \<subseteq> ball p e1"
             using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto
           ultimately show ?thesis using \<open>e2>0\<close> by auto
         qed
@@ -3303,15 +3303,15 @@
           define zp where "zp \<equiv> zer_poly f p"
           def f'\<equiv>"\<lambda>w. zp w * (w - p) ^ zo"
           def ff'\<equiv>"(\<lambda>x. deriv f' x * h x / f' x)"
-          have "f holomorphic_on ball p e1" 
+          have "f holomorphic_on ball p e1"
             proof -
-              have "ball p e1 \<subseteq> s - poles" 
+              have "ball p e1 \<subseteq> s - poles"
                 using avoid_def ball_subset_cball e1_avoid that zeros_def by fastforce
               thus ?thesis using f_holo by auto
             qed
-          moreover have "f p = 0" using \<open>p\<in>zeros\<close> 
+          moreover have "f p = 0" using \<open>p\<in>zeros\<close>
             using DiffD1 mem_Collect_eq zeros_def by blast
-          moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0" 
+          moreover have "\<exists>w\<in>ball p e1. f w \<noteq> 0"
             proof -
               def p'\<equiv>"p+e1/2"
               have "p'\<in>ball p e1" and "p'\<noteq>p" using \<open>e1>0\<close> unfolding p'_def by (auto simp add:dist_norm)
@@ -3319,24 +3319,24 @@
                 apply (rule_tac x=p' in bexI)
                 by (auto simp add:zeros_def)
             qed
-          ultimately obtain r where 
+          ultimately obtain r where
               "0 < zo" "r>0"
               "cball p r \<subseteq> ball p e1" and
               pp_holo:"zp holomorphic_on cball p r" and
               pp_po:"(\<forall>w\<in>cball p r. f w = zp w * (w - p) ^ zo \<and> zp w \<noteq> 0)"
             using zorder_exist[of "ball p e1" p f,simplified,OF \<open>e1>0\<close>] unfolding zo_def zp_def
-            by auto    
+            by auto
           define e2 where "e2 \<equiv> r/2"
           have "e2>0" using \<open>r>0\<close> unfolding e2_def by auto
-          define anal where "anal \<equiv> \<lambda>w. deriv zp w * h w / zp w" 
+          define anal where "anal \<equiv> \<lambda>w. deriv zp w * h w / zp w"
           define prin where "prin \<equiv> \<lambda>w. of_nat zo * h w / (w - p)"
           have "((\<lambda>w.  prin w + anal w) has_contour_integral c * zo * h p) (circlepath p e2)"
             proof (rule  has_contour_integral_add[of _ _ _ _ 0,simplified])
-              have "ball p r \<subseteq> s" 
+              have "ball p r \<subseteq> s"
                 using \<open>cball p r \<subseteq> ball p e1\<close> avoid_def ball_subset_cball e1_avoid by blast
               then have "cball p e2 \<subseteq> s"
-                using \<open>r>0\<close> unfolding e2_def by auto  
-              then have "(\<lambda>w. of_nat zo * h w) holomorphic_on cball p e2" 
+                using \<open>r>0\<close> unfolding e2_def by auto
+              then have "(\<lambda>w. of_nat zo * h w) holomorphic_on cball p e2"
                 using h_holo
                 by (auto intro!: holomorphic_intros)
               then show "(prin has_contour_integral c * of_nat zo * h p ) (circlepath p e2)"
@@ -3344,7 +3344,7 @@
                   \<open>e2>0\<close>
                 unfolding prin_def
                 by (auto simp add: mult.assoc)
-              have "anal holomorphic_on ball p r" unfolding anal_def 
+              have "anal holomorphic_on ball p r" unfolding anal_def
                 using pp_holo h_holo pp_po \<open>ball p r \<subseteq> s\<close>
                 by (auto intro!: holomorphic_intros)
               then show "(anal has_contour_integral 0) (circlepath p e2)"
@@ -3353,19 +3353,19 @@
             qed
           then have "cont_zero ff' p e2" unfolding cont_zero_def zo_def
             proof (elim has_contour_integral_eq)
-              fix w assume "w \<in> path_image (circlepath p e2)" 
+              fix w assume "w \<in> path_image (circlepath p e2)"
               then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
               define wp where "wp \<equiv> w-p"
-              have "wp\<noteq>0" and "zp w \<noteq>0" 
-                unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto 
-              moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo" 
+              have "wp\<noteq>0" and "zp w \<noteq>0"
+                unfolding wp_def using \<open>w\<noteq>p\<close> \<open>w\<in>ball p r\<close> pp_po by auto
+              moreover have der_f':"deriv f' w = zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo"
                 proof (rule DERIV_imp_deriv)
                   define der where "der \<equiv> zo * zp w * (w-p)^(zo-1) + deriv zp w * (w-p)^zo"
                   have po:"zo = Suc (zo - Suc 0) " using \<open>zo>0\<close> by auto
-                  have "(zp has_field_derivative (deriv zp w)) (at w)" 
+                  have "(zp has_field_derivative (deriv zp w)) (at w)"
                     using DERIV_deriv_iff_has_field_derivative pp_holo
                     by (meson Topology_Euclidean_Space.open_ball \<open>w \<in> ball p r\<close> ball_subset_cball holomorphic_derivI holomorphic_on_subset)
-                  then show "(f' has_field_derivative  der) (at w)" 
+                  then show "(f' has_field_derivative  der) (at w)"
                     using \<open>w\<noteq>p\<close> \<open>zo>0\<close> unfolding der_def f'_def
                     by (auto intro!: derivative_eq_intros simp add:field_simps)
                 qed
@@ -3377,16 +3377,16 @@
                 apply (auto simp add:field_simps)
                 by (metis Suc_diff_Suc minus_nat.diff_0 power_Suc)
             qed
-          then have "cont_zero ff p e2" unfolding cont_zero_def   
+          then have "cont_zero ff p e2" unfolding cont_zero_def
             proof (elim has_contour_integral_eq)
               fix w assume "w \<in> path_image (circlepath p e2)"
               then have "w\<in>ball p r" and "w\<noteq>p" unfolding e2_def using \<open>r>0\<close> by auto
-              have "deriv f' w =  deriv f w" 
+              have "deriv f' w =  deriv f w"
                 proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"])
                   show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo
                     by (auto intro!: holomorphic_intros)
-                next                
-                  have "ball p e1 - {p} \<subseteq> s - poles" 
+                next
+                  have "ball p e1 - {p} \<subseteq> s - poles"
                     using avoid_def ball_subset_cball e1_avoid by auto
                   then have "ball p r - {p} \<subseteq> s - poles" using \<open>cball p r \<subseteq> ball p e1\<close>
                     using ball_subset_cball by blast
@@ -3398,37 +3398,37 @@
                   show "w \<in> ball p r - {p}" using \<open>w\<in>ball p r\<close> \<open>w\<noteq>p\<close> by auto
                 next
                   fix x assume "x \<in> ball p r - {p}"
-                  then show "f' x = f x" 
+                  then show "f' x = f x"
                     using pp_po unfolding f'_def by auto
                 qed
-              moreover have " f' w  =  f w " 
+              moreover have " f' w  =  f w "
                 using \<open>w \<in> ball p r\<close> ball_subset_cball subset_iff pp_po unfolding f'_def by auto
               ultimately show "ff' w = ff w"
                 unfolding ff'_def ff_def by simp
             qed
-          moreover have "cball p e2 \<subseteq> ball p e1" 
+          moreover have "cball p e2 \<subseteq> ball p e1"
             using \<open>0 < r\<close> \<open>cball p r \<subseteq> ball p e1\<close> e2_def by auto
           ultimately show ?thesis using \<open>e2>0\<close> by auto
         qed
       then obtain e3 where e3:"p\<in>zeros \<longrightarrow> e3>0 \<and> cball p e3 \<subseteq> ball p e1 \<and> cont_zero ff p e3"
-        by auto          
+        by auto
       define e4 where "e4 \<equiv> if p\<in>poles then e2 else if p\<in>zeros then e3 else e1"
       have "e4>0" using e2 e3 \<open>e1>0\<close> unfolding e4_def by auto
       moreover have "avoid p e4" using e2 e3 \<open>e1>0\<close> e1_avoid unfolding e4_def avoid_def by auto
-      moreover have "p\<in>poles \<longrightarrow> cont_pole ff p e4" and "p\<in>zeros \<longrightarrow> cont_zero ff p e4" 
+      moreover have "p\<in>poles \<longrightarrow> cont_pole ff p e4" and "p\<in>zeros \<longrightarrow> cont_zero ff p e4"
         by (auto simp add: e2 e3 e4_def zeros_def)
       ultimately show ?thesis by auto
     qed
-  then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p) 
+  then obtain get_e where get_e:"\<forall>p\<in>s. get_e p>0 \<and> avoid p (get_e p)
       \<and> (p\<in>poles \<longrightarrow> cont_pole ff p (get_e p)) \<and> (p\<in>zeros \<longrightarrow> cont_zero ff p (get_e p))"
     by metis
   define cont where "cont \<equiv> \<lambda>p. contour_integral (circlepath p (get_e p)) ff"
   define w where "w \<equiv> \<lambda>p. winding_number g p"
   have "contour_integral g ff = (\<Sum>p\<in>zeros \<union> poles. w p * cont p)"
     unfolding cont_def w_def
-    proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop 
+    proof (rule Cauchy_theorem_singularities[OF \<open>open s\<close> \<open>connected s\<close> finite _ \<open>valid_path g\<close> loop
         path_img homo])
-      have "open (s - (zeros \<union> poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto 
+      have "open (s - (zeros \<union> poles))" using open_Diff[OF _ finite_imp_closed[OF finite]] \<open>open s\<close> by auto
       then show "ff holomorphic_on s - (zeros \<union> poles)" unfolding ff_def using f_holo h_holo
         by (auto intro!: holomorphic_intros simp add:zeros_def)
     next
@@ -3436,12 +3436,12 @@
         using get_e using avoid_def by blast
     qed
   also have "... = (\<Sum>p\<in>zeros. w p * cont p) + (\<Sum>p\<in>poles. w p * cont p)"
-    using finite 
+    using finite
     apply (subst setsum.union_disjoint)
     by (auto simp add:zeros_def)
   also have "... = c * ((\<Sum>p\<in>zeros. w p *  h p * zorder f p) - (\<Sum>p\<in>poles. w p *  h p * porder f p))"
     proof -
-      have "(\<Sum>p\<in>zeros. w p * cont p) = (\<Sum>p\<in>zeros. c * w p *  h p * zorder f p)" 
+      have "(\<Sum>p\<in>zeros. w p * cont p) = (\<Sum>p\<in>zeros. c * w p *  h p * zorder f p)"
         proof (rule setsum.cong[of zeros zeros,simplified])
           fix p assume "p \<in> zeros"
           show "w p * cont p = c * w p * h p * (zorder f p)"
@@ -3449,7 +3449,7 @@
               assume "p \<in> s"
               have "cont p = c * h p * (zorder f p)" unfolding cont_def
                 apply (rule contour_integral_unique)
-                using get_e \<open>p\<in>s\<close> \<open>p\<in>zeros\<close> unfolding cont_zero_def                
+                using get_e \<open>p\<in>s\<close> \<open>p\<in>zeros\<close> unfolding cont_zero_def
                 by (metis mult.assoc mult.commute)
               thus ?thesis by auto
             next
@@ -3458,10 +3458,10 @@
               then show ?thesis by auto
             qed
         qed
-      then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros.  w p *  h p * zorder f p)" 
+      then have "(\<Sum>p\<in>zeros. w p * cont p) = c * (\<Sum>p\<in>zeros.  w p *  h p * zorder f p)"
         apply (subst setsum_right_distrib)
         by (simp add:algebra_simps)
-      moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles.  - c * w p *  h p * porder f p)" 
+      moreover have "(\<Sum>p\<in>poles. w p * cont p) = (\<Sum>p\<in>poles.  - c * w p *  h p * porder f p)"
         proof (rule setsum.cong[of poles poles,simplified])
           fix p assume "p \<in> poles"
           show "w p * cont p = - c * w p * h p * (porder f p)"
@@ -3469,7 +3469,7 @@
               assume "p \<in> s"
               have "cont p = - c * h p * (porder f p)" unfolding cont_def
                 apply (rule contour_integral_unique)
-                using get_e \<open>p\<in>s\<close> \<open>p\<in>poles\<close> unfolding cont_pole_def               
+                using get_e \<open>p\<in>s\<close> \<open>p\<in>poles\<close> unfolding cont_pole_def
                 by (metis mult.assoc mult.commute)
               thus ?thesis by auto
             next
@@ -3478,11 +3478,11 @@
               then show ?thesis by auto
             qed
         qed
-      then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p *  h p * porder f p)" 
+      then have "(\<Sum>p\<in>poles. w p * cont p) = - c * (\<Sum>p\<in>poles. w p *  h p * porder f p)"
         apply (subst setsum_right_distrib)
         by (simp add:algebra_simps)
       ultimately show ?thesis by (simp add: right_diff_distrib)
-    qed      
+    qed
   finally show ?thesis unfolding w_def ff_def c_def by auto
 qed
 
@@ -3493,59 +3493,59 @@
   defines "fg\<equiv>(\<lambda>p. f p+ g p)"
   defines "zeros_fg\<equiv>{p. fg p =0}" and "zeros_f\<equiv>{p. f p=0}"
   assumes
-    "open s" and "connected s" and 
-    "finite zeros_fg" and 
+    "open s" and "connected s" and
+    "finite zeros_fg" and
     "finite zeros_f" and
-    f_holo:"f holomorphic_on s" and 
+    f_holo:"f holomorphic_on s" and
     g_holo:"g holomorphic_on s" and
-    "valid_path \<gamma>" and 
+    "valid_path \<gamma>" and
     loop:"pathfinish \<gamma> = pathstart \<gamma>" and
     path_img:"path_image \<gamma> \<subseteq> s " and
-    path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and 
+    path_less:"\<forall>z\<in>path_image \<gamma>. cmod(f z) > cmod(g z)" and
     homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number \<gamma> z = 0"
-  shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p) 
+  shows "(\<Sum>p\<in>zeros_fg. winding_number \<gamma> p * zorder fg p)
           = (\<Sum>p\<in>zeros_f. winding_number \<gamma> p * zorder f p)"
 proof -
   have path_fg:"path_image \<gamma> \<subseteq> s - zeros_fg"
     proof -
-      have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z 
+      have False when "z\<in>path_image \<gamma>" and "f z + g z=0" for z
         proof -
           have "cmod (f z) > cmod (g z)" using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
           moreover have "f z = - g z"  using \<open>f z + g z =0\<close> by (simp add: eq_neg_iff_add_eq_0)
           then have "cmod (f z) = cmod (g z)" by auto
           ultimately show False by auto
         qed
-      then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto 
+      then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto
     qed
   have path_f:"path_image \<gamma> \<subseteq> s - zeros_f"
     proof -
-      have False when "z\<in>path_image \<gamma>" and "f z =0" for z 
+      have False when "z\<in>path_image \<gamma>" and "f z =0" for z
         proof -
           have "cmod (g z) < cmod (f z) " using \<open>z\<in>path_image \<gamma>\<close> path_less by auto
-          then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto          
+          then have "cmod (g z) < 0" using \<open>f z=0\<close> by auto
           then show False by auto
         qed
-      then show ?thesis unfolding zeros_f_def using path_img by auto 
+      then show ?thesis unfolding zeros_f_def using path_img by auto
     qed
   define w where "w \<equiv> \<lambda>p. winding_number \<gamma> p"
   define c where "c \<equiv> 2 * complex_of_real pi * \<i>"
-  define h where "h \<equiv> \<lambda>p. g p / f p + 1" 
+  define h where "h \<equiv> \<lambda>p. g p / f p + 1"
   obtain spikes
     where "finite spikes" and spikes: "\<forall>x\<in>{0..1} - spikes. \<gamma> differentiable at x"
     using \<open>valid_path \<gamma>\<close>
     by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq)
-  have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>" 
+  have h_contour:"((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
     proof -
-      have outside_img:"0 \<in> outside (path_image (h o \<gamma>))" 
+      have outside_img:"0 \<in> outside (path_image (h o \<gamma>))"
         proof -
           have "h p \<in> ball 1 1" when "p\<in>path_image \<gamma>" for p
             proof -
-              have "cmod (g p/f p) <1" using path_less[rule_format,OF that] 
+              have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
                 apply (cases "cmod (f p) = 0")
                 by (auto simp add: norm_divide)
               then show ?thesis unfolding h_def by (auto simp add:dist_complex_def)
             qed
-          then have "path_image (h o \<gamma>) \<subseteq> ball 1 1" 
+          then have "path_image (h o \<gamma>) \<subseteq> ball 1 1"
             by (simp add: image_subset_iff path_image_compose)
           moreover have " (0::complex) \<notin> ball 1 1" by (simp add: dist_norm)
           ultimately show "?thesis"
@@ -3553,67 +3553,67 @@
         qed
       have valid_h:"valid_path (h \<circ> \<gamma>)"
         proof (rule valid_path_compose_holomorphic[OF \<open>valid_path \<gamma>\<close> _ _ path_f])
-          show "h holomorphic_on s - zeros_f" 
+          show "h holomorphic_on s - zeros_f"
             unfolding h_def using f_holo g_holo
             by (auto intro!: holomorphic_intros simp add:zeros_f_def)
         next
           show "open (s - zeros_f)" using \<open>finite zeros_f\<close> \<open>open s\<close> finite_imp_closed
             by auto
         qed
-      have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)" 
+      have "((\<lambda>z. 1/z) has_contour_integral 0) (h \<circ> \<gamma>)"
         proof -
           have "0 \<notin> path_image (h \<circ> \<gamma>)" using outside_img by (simp add: outside_def)
           then have "((\<lambda>z. 1/z) has_contour_integral c * winding_number (h \<circ> \<gamma>) 0) (h \<circ> \<gamma>)"
-            using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h 
+            using has_contour_integral_winding_number[of "h o \<gamma>" 0,simplified] valid_h
             unfolding c_def by auto
-          moreover have "winding_number (h o \<gamma>) 0 = 0" 
+          moreover have "winding_number (h o \<gamma>) 0 = 0"
             proof -
               have "0 \<in> outside (path_image (h \<circ> \<gamma>))" using outside_img .
-              moreover have "path (h o \<gamma>)" 
+              moreover have "path (h o \<gamma>)"
                 using valid_h  by (simp add: valid_path_imp_path)
-              moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)" 
+              moreover have "pathfinish (h o \<gamma>) = pathstart (h o \<gamma>)"
                 by (simp add: loop pathfinish_compose pathstart_compose)
               ultimately show ?thesis using winding_number_zero_in_outside by auto
             qed
           ultimately show ?thesis by auto
         qed
       moreover have "vector_derivative (h \<circ> \<gamma>) (at x) = vector_derivative \<gamma> (at x) * deriv h (\<gamma> x)"
-          when "x\<in>{0..1} - spikes" for x 
+          when "x\<in>{0..1} - spikes" for x
         proof (rule vector_derivative_chain_at_general)
           show "\<gamma> differentiable at x" using that \<open>valid_path \<gamma>\<close> spikes by auto
         next
           define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
           define t where "t \<equiv> \<gamma> x"
-          have "f t\<noteq>0" unfolding zeros_f_def t_def 
+          have "f t\<noteq>0" unfolding zeros_f_def t_def
             by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that)
-          moreover have "t\<in>s" 
+          moreover have "t\<in>s"
             using contra_subsetD path_image_def path_fg t_def that by fastforce
           ultimately have "(h has_field_derivative der t) (at t)"
             unfolding h_def der_def using g_holo f_holo \<open>open s\<close>
             by (auto intro!: holomorphic_derivI derivative_eq_intros )
           then show "\<exists>g'. (h has_field_derivative g') (at (\<gamma> x))" unfolding t_def by auto
         qed
-      then have " (op / 1 has_contour_integral 0) (h \<circ> \<gamma>) 
-          = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"  
+      then have " (op / 1 has_contour_integral 0) (h \<circ> \<gamma>)
+          = ((\<lambda>x. deriv h x / h x) has_contour_integral 0) \<gamma>"
         unfolding has_contour_integral
         apply (intro has_integral_spike_eq[OF negligible_finite, OF \<open>finite spikes\<close>])
         by auto
       ultimately show ?thesis by auto
     qed
-  then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0" 
+  then have "contour_integral \<gamma> (\<lambda>x. deriv h x / h x) = 0"
     using  contour_integral_unique by simp
-  moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x) 
-      + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)" 
+  moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = contour_integral \<gamma> (\<lambda>x. deriv f x / f x)
+      + contour_integral \<gamma> (\<lambda>p. deriv h p / h p)"
     proof -
-      have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>" 
+      have "(\<lambda>p. deriv f p / f p) contour_integrable_on \<gamma>"
         proof (rule contour_integrable_holomorphic_simple[OF _ _ \<open>valid_path \<gamma>\<close> path_f])
-          show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close> 
+          show "open (s - zeros_f)" using finite_imp_closed[OF \<open>finite zeros_f\<close>] \<open>open s\<close>
             by auto
           then show "(\<lambda>p. deriv f p / f p) holomorphic_on s - zeros_f"
             using f_holo
             by (auto intro!: holomorphic_intros simp add:zeros_f_def)
         qed
-      moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>" 
+      moreover have "(\<lambda>p. deriv h p / h p) contour_integrable_on \<gamma>"
         using h_contour
         by (simp add: has_contour_integral_integrable)
       ultimately have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x + deriv h x / h x) =
@@ -3621,16 +3621,16 @@
         using contour_integral_add[of "(\<lambda>p. deriv f p / f p)" \<gamma> "(\<lambda>p. deriv h p / h p)" ]
         by auto
       moreover have "deriv fg p / fg p =  deriv f p / f p + deriv h p / h p"
-          when "p\<in> path_image \<gamma>" for p 
+          when "p\<in> path_image \<gamma>" for p
         proof -
           have "fg p\<noteq>0" and "f p\<noteq>0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def
             by auto
-          have "h p\<noteq>0" 
+          have "h p\<noteq>0"
             proof (rule ccontr)
               assume "\<not> h p \<noteq> 0"
               then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2)
               then have "cmod (g p/f p) = 1" by auto
-              moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] 
+              moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that]
                 apply (cases "cmod (f p) = 0")
                 by (auto simp add: norm_divide)
               ultimately show False by auto
@@ -3640,7 +3640,7 @@
             by auto
           have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)"
             proof -
-              define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"          
+              define der where "der \<equiv> \<lambda>p. (deriv g p * f p - g p * deriv f p)/(f p * f p)"
               have "p\<in>s" using path_img that by auto
               then have "(h has_field_derivative der p) (at p)"
                 unfolding h_def der_def using g_holo f_holo \<open>open s\<close> \<open>f p\<noteq>0\<close>
@@ -3648,17 +3648,17 @@
               then show ?thesis unfolding der_def using DERIV_imp_deriv by auto
             qed
           show ?thesis
-            apply (simp only:der_fg der_h)   
+            apply (simp only:der_fg der_h)
             apply (auto simp add:field_simps \<open>h p\<noteq>0\<close> \<open>f p\<noteq>0\<close> \<open>fg p\<noteq>0\<close>)
-            by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def) 
+            by (auto simp add:field_simps h_def \<open>f p\<noteq>0\<close> fg_def)
         qed
-      then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p) 
+      then have "contour_integral \<gamma> (\<lambda>p. deriv fg p / fg p)
           = contour_integral \<gamma> (\<lambda>p. deriv f p / f p + deriv h p / h p)"
-        by (elim contour_integral_eq)        
+        by (elim contour_integral_eq)
       ultimately show ?thesis by auto
     qed
   moreover have "contour_integral \<gamma> (\<lambda>x. deriv fg x / fg x) = c * (\<Sum>p\<in>zeros_fg. w p * zorder fg p)"
-    unfolding c_def zeros_fg_def w_def 
+    unfolding c_def zeros_fg_def w_def
     proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
         , of _ "{}" "\<lambda>_. 1",simplified])
       show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto
@@ -3666,13 +3666,13 @@
       show " finite {p. fg p = 0}" using \<open>finite zeros_fg\<close> unfolding zeros_fg_def .
     qed
   moreover have "contour_integral \<gamma> (\<lambda>x. deriv f x / f x) = c * (\<Sum>p\<in>zeros_f. w p * zorder f p)"
-    unfolding c_def zeros_f_def w_def 
+    unfolding c_def zeros_f_def w_def
     proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close> _ _ \<open>valid_path \<gamma>\<close> loop _ homo
         , of _ "{}" "\<lambda>_. 1",simplified])
       show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto
       show "path_image \<gamma> \<subseteq> s - {p. f p = 0}" using path_f unfolding zeros_f_def .
       show " finite {p. f p = 0}" using \<open>finite zeros_f\<close> unfolding zeros_f_def .
-    qed  
+    qed
   ultimately have " c* (\<Sum>p\<in>zeros_fg. w p * (zorder fg p)) = c* (\<Sum>p\<in>zeros_f. w p * (zorder f p))"
     by auto
   then show ?thesis unfolding c_def using w_def by auto
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Continuous_Extension.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,547 @@
+(*  Title:      HOL/Multivariate_Analysis/Extension.thy
+    Authors:    LC Paulson, based on material from HOL Light
+*)
+
+section \<open>Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\<close>
+
+theory Continuous_Extension
+imports Convex_Euclidean_Space
+begin
+
+subsection\<open>Partitions of unity subordinate to locally finite open coverings\<close>
+
+text\<open>A difference from HOL Light: all summations over infinite sets equal zero,
+   so the "support" must be made explicit in the summation below!\<close>
+
+proposition subordinate_partition_of_unity:
+  fixes S :: "'a :: euclidean_space set"
+  assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
+      and fin: "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
+  obtains F :: "['a set, 'a] \<Rightarrow> real"
+    where "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x \<in> S. 0 \<le> F U x)"
+      and "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
+      and "\<And>x. x \<in> S \<Longrightarrow> supp_setsum (\<lambda>W. F W x) \<C> = 1"
+      and "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}"
+proof (cases "\<exists>W. W \<in> \<C> \<and> S \<subseteq> W")
+  case True
+    then obtain W where "W \<in> \<C>" "S \<subseteq> W" by metis
+    then show ?thesis
+      apply (rule_tac F = "\<lambda>V x. if V = W then 1 else 0" in that)
+      apply (auto simp: continuous_on_const supp_setsum_def support_on_def)
+      done
+next
+  case False
+    have nonneg: "0 \<le> supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" for x
+      by (simp add: supp_setsum_def setsum_nonneg)
+    have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> V" for V x
+    proof -
+      have "closedin (subtopology euclidean S) (S - V)"
+        by (simp add: Diff_Diff_Int Diff_subset closedin_def opC openin_open_Int \<open>V \<in> \<C>\<close>)
+      with that False setdist_eq_0_closedin [of S "S-V" x] setdist_pos_le [of "{x}" "S - V"]
+        show ?thesis
+          by (simp add: order_class.order.order_iff_strict)
+    qed
+    have ss_pos: "0 < supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x
+    proof -
+      obtain U where "U \<in> \<C>" "x \<in> U" using \<open>x \<in> S\<close> \<open>S \<subseteq> \<Union>\<C>\<close>
+        by blast
+      obtain V where "open V" "x \<in> V" "finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
+        using \<open>x \<in> S\<close> fin by blast
+      then have *: "finite {A \<in> \<C>. \<not> S \<subseteq> A \<and> x \<notin> closure (S - A)}"
+        using closure_def that by (blast intro: rev_finite_subset)
+      have "x \<notin> closure (S - U)"
+        by (metis \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> less_irrefl sd_pos setdist_eq_0_sing_1 that)
+      then show ?thesis
+        apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def)
+        apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U])
+        using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> False
+        apply (auto simp: setdist_pos_le sd_pos that)
+        done
+    qed
+    define F where
+      "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>
+                 else 0"
+    show ?thesis
+    proof (rule_tac F = F in that)
+      have "continuous_on S (F U)" if "U \<in> \<C>" for U
+      proof -
+        have *: "continuous_on S (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
+        proof (clarsimp simp add: continuous_on_eq_continuous_within)
+          fix x assume "x \<in> S"
+          then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}"
+            using assms by blast
+          then have OSX: "openin (subtopology euclidean S) (S \<inter> X)" by blast
+          have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow>
+                     (\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V))
+                     = supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>"
+            apply (simp add: supp_setsum_def)
+            apply (rule setsum.mono_neutral_right [OF finX])
+            apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff)
+            apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE)
+            done
+          show "continuous (at x within S) (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
+            apply (rule continuous_transform_within_openin
+                     [where f = "\<lambda>x. (setsum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})"
+                        and S ="S \<inter> X"])
+            apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+
+            apply (simp add: sumeq)
+            done
+        qed
+        show ?thesis
+          apply (simp add: F_def)
+          apply (rule continuous_intros *)+
+          using ss_pos apply force
+          done
+      qed
+      moreover have "\<lbrakk>U \<in> \<C>; x \<in> S\<rbrakk> \<Longrightarrow> 0 \<le> F U x" for U x
+        using nonneg [of x] by (simp add: F_def divide_simps setdist_pos_le)
+      ultimately show "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x\<in>S. 0 \<le> F U x)"
+        by metis
+    next
+      show "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
+        by (simp add: setdist_eq_0_sing_1 closure_def F_def)
+    next
+      show "supp_setsum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x
+        using that ss_pos [OF that]
+        by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric])
+    next
+      show "\<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" if "x \<in> S" for x
+        using fin [OF that] that
+        by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset)
+    qed
+qed
+
+
+subsection\<open>Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\<close>
+
+lemma Urysohn_both_ne:
+  assumes US: "closedin (subtopology euclidean U) S"
+      and UT: "closedin (subtopology euclidean U) T"
+      and "S \<inter> T = {}" "S \<noteq> {}" "T \<noteq> {}" "a \<noteq> b"
+  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+    where "continuous_on U f"
+          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
+          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
+          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
+proof -
+  have S0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} S = 0 \<longleftrightarrow> x \<in> S"
+    using \<open>S \<noteq> {}\<close>  US setdist_eq_0_closedin  by auto
+  have T0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} T = 0 \<longleftrightarrow> x \<in> T"
+    using \<open>T \<noteq> {}\<close>  UT setdist_eq_0_closedin  by auto
+  have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \<in> U" for x
+  proof -
+    have "~ (setdist {x} S = 0 \<and> setdist {x} T = 0)"
+      using assms by (metis IntI empty_iff setdist_eq_0_closedin that)
+    then show ?thesis
+      by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le)
+  qed
+  define f where "f \<equiv> \<lambda>x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)"
+  show ?thesis
+  proof (rule_tac f = f in that)
+    show "continuous_on U f"
+      using sdpos unfolding f_def
+      by (intro continuous_intros | force)+
+    show "f x \<in> closed_segment a b" if "x \<in> U" for x
+        unfolding f_def
+      apply (simp add: closed_segment_def)
+      apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI)
+      using sdpos that apply (simp add: algebra_simps)
+      done
+    show "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
+      using S0 \<open>a \<noteq> b\<close> f_def sdpos by force
+    show "(f x = b \<longleftrightarrow> x \<in> T)" if "x \<in> U" for x
+    proof -
+      have "f x = b \<longleftrightarrow> (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1"
+        unfolding f_def
+        apply (rule iffI)
+         apply (metis  \<open>a \<noteq> b\<close> add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force)
+        done
+      also have "...  \<longleftrightarrow> setdist {x} T = 0 \<and> setdist {x} S \<noteq> 0"
+        using sdpos that
+        by (simp add: divide_simps) linarith
+      also have "...  \<longleftrightarrow> x \<in> T"
+        using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>S \<inter> T = {}\<close> that
+        by (force simp: S0 T0)
+      finally show ?thesis .
+    qed
+  qed
+qed
+
+proposition Urysohn_local_strong:
+  assumes US: "closedin (subtopology euclidean U) S"
+      and UT: "closedin (subtopology euclidean U) T"
+      and "S \<inter> T = {}" "a \<noteq> b"
+  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+    where "continuous_on U f"
+          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
+          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
+          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
+proof (cases "S = {}")
+  case True show ?thesis
+  proof (cases "T = {}")
+    case True show ?thesis
+    proof (rule_tac f = "\<lambda>x. midpoint a b" in that)
+      show "continuous_on U (\<lambda>x. midpoint a b)"
+        by (intro continuous_intros)
+      show "midpoint a b \<in> closed_segment a b"
+        using csegment_midpoint_subset by blast
+      show "(midpoint a b = a) = (x \<in> S)" for x
+        using \<open>S = {}\<close> \<open>a \<noteq> b\<close> by simp
+      show "(midpoint a b = b) = (x \<in> T)" for x
+        using \<open>T = {}\<close> \<open>a \<noteq> b\<close> by simp
+    qed
+  next
+    case False
+    show ?thesis
+    proof (cases "T = U")
+      case True with \<open>S = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
+        by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const)
+    next
+      case False
+      with UT closedin_subset obtain c where c: "c \<in> U" "c \<notin> T"
+        by fastforce
+      obtain f where f: "continuous_on U f"
+                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment (midpoint a b) b"
+                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
+                "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
+        apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"])
+        using c \<open>T \<noteq> {}\<close> assms apply simp_all
+        done
+      show ?thesis
+        apply (rule_tac f=f in that)
+        using \<open>S = {}\<close> \<open>T \<noteq> {}\<close> f csegment_midpoint_subset notin_segment_midpoint [OF \<open>a \<noteq> b\<close>]
+        apply force+
+        done
+    qed
+  qed
+next
+  case False
+  show ?thesis
+  proof (cases "T = {}")
+    case True show ?thesis
+    proof (cases "S = U")
+      case True with \<open>T = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
+        by (rule_tac f = "\<lambda>x. a" in that) (auto simp: continuous_on_const)
+    next
+      case False
+      with US closedin_subset obtain c where c: "c \<in> U" "c \<notin> S"
+        by fastforce
+      obtain f where f: "continuous_on U f"
+                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a (midpoint a b)"
+                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
+                "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
+        apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"])
+        using c \<open>S \<noteq> {}\<close> assms apply simp_all
+        apply (metis midpoint_eq_endpoint)
+        done
+      show ?thesis
+        apply (rule_tac f=f in that)
+        using \<open>S \<noteq> {}\<close> \<open>T = {}\<close> f  \<open>a \<noteq> b\<close>
+        apply simp_all
+        apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff)
+        apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint)
+        done
+    qed
+  next
+    case False
+    show ?thesis
+      using Urysohn_both_ne [OF US UT \<open>S \<inter> T = {}\<close> \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>a \<noteq> b\<close>] that
+      by blast
+  qed
+qed
+
+lemma Urysohn_local:
+  assumes US: "closedin (subtopology euclidean U) S"
+      and UT: "closedin (subtopology euclidean U) T"
+      and "S \<inter> T = {}"
+  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+    where "continuous_on U f"
+          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
+          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
+          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
+proof (cases "a = b")
+  case True then show ?thesis
+    by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const)
+next
+  case False
+  then show ?thesis
+    apply (rule Urysohn_local_strong [OF assms])
+    apply (erule that, assumption)
+    apply (meson US closedin_singleton closedin_trans)
+    apply (meson UT closedin_singleton closedin_trans)
+    done
+qed
+
+lemma Urysohn_strong:
+  assumes US: "closed S"
+      and UT: "closed T"
+      and "S \<inter> T = {}" "a \<noteq> b"
+  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+    where "continuous_on UNIV f"
+          "\<And>x. f x \<in> closed_segment a b"
+          "\<And>x. f x = a \<longleftrightarrow> x \<in> S"
+          "\<And>x. f x = b \<longleftrightarrow> x \<in> T"
+apply (rule Urysohn_local_strong [of UNIV S T])
+using assms
+apply (auto simp: closed_closedin)
+done
+
+proposition Urysohn:
+  assumes US: "closed S"
+      and UT: "closed T"
+      and "S \<inter> T = {}"
+  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+    where "continuous_on UNIV f"
+          "\<And>x. f x \<in> closed_segment a b"
+          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
+          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
+apply (rule Urysohn_local [of UNIV S T a b])
+using assms
+apply (auto simp: closed_closedin)
+done
+
+
+subsection\<open> The Dugundji extension theorem, and Tietze variants as corollaries.\<close>
+
+text\<open>J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367.
+http://projecteuclid.org/euclid.pjm/1103052106\<close>
+
+theorem Dugundji:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
+  assumes "convex C" "C \<noteq> {}"
+      and cloin: "closedin (subtopology euclidean U) S"
+      and contf: "continuous_on S f" and "f ` S \<subseteq> C"
+  obtains g where "continuous_on U g" "g ` U \<subseteq> C"
+                  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+proof (cases "S = {}")
+  case True then show thesis
+    apply (rule_tac g="\<lambda>x. @y. y \<in> C" in that)
+      apply (rule continuous_intros)
+     apply (meson all_not_in_conv \<open>C \<noteq> {}\<close> image_subsetI someI_ex, simp)
+    done
+next
+  case False
+  then have sd_pos: "\<And>x. \<lbrakk>x \<in> U; x \<notin> S\<rbrakk> \<Longrightarrow> 0 < setdist {x} S"
+    using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce
+  define \<B> where "\<B> = {ball x (setdist {x} S / 2) |x. x \<in> U - S}"
+  have [simp]: "\<And>T. T \<in> \<B> \<Longrightarrow> open T"
+    by (auto simp: \<B>_def)
+  have USS: "U - S \<subseteq> \<Union>\<B>"
+    by (auto simp: sd_pos \<B>_def)
+  obtain \<C> where USsub: "U - S \<subseteq> \<Union>\<C>"
+       and nbrhd: "\<And>U. U \<in> \<C> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<B> \<and> U \<subseteq> T)"
+       and fin: "\<And>x. x \<in> U - S
+                     \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> \<C> \<and> U \<inter> V \<noteq> {}}"
+    using paracompact [OF USS] by auto
+  have "\<exists>v a. v \<in> U \<and> v \<notin> S \<and> a \<in> S \<and>
+              T \<subseteq> ball v (setdist {v} S / 2) \<and>
+              dist v a \<le> 2 * setdist {v} S" if "T \<in> \<C>" for T
+  proof -
+    obtain v where v: "T \<subseteq> ball v (setdist {v} S / 2)" "v \<in> U" "v \<notin> S"
+      using \<open>T \<in> \<C>\<close> nbrhd by (force simp: \<B>_def)
+    then obtain a where "a \<in> S" "dist v a < 2 * setdist {v} S"
+      using setdist_ltE [of "{v}" S "2 * setdist {v} S"]
+      using False sd_pos by force
+    with v show ?thesis
+      apply (rule_tac x=v in exI)
+      apply (rule_tac x=a in exI, auto)
+      done
+  qed
+  then obtain \<V> \<A> where
+    VA: "\<And>T. T \<in> \<C> \<Longrightarrow> \<V> T \<in> U \<and> \<V> T \<notin> S \<and> \<A> T \<in> S \<and>
+              T \<subseteq> ball (\<V> T) (setdist {\<V> T} S / 2) \<and>
+              dist (\<V> T) (\<A> T) \<le> 2 * setdist {\<V> T} S"
+    by metis
+  have sdle: "setdist {\<V> T} S \<le> 2 * setdist {v} S" if "T \<in> \<C>" "v \<in> T" for T v
+    using setdist_Lipschitz [of "\<V> T" S v] VA [OF \<open>T \<in> \<C>\<close>] \<open>v \<in> T\<close> by auto
+  have d6: "dist a (\<A> T) \<le> 6 * dist a v" if "T \<in> \<C>" "v \<in> T" "a \<in> S" for T v a
+  proof -
+    have "dist (\<V> T) v < setdist {\<V> T} S / 2"
+      using that VA mem_ball by blast
+    also have "... \<le> setdist {v} S"
+      using sdle [OF \<open>T \<in> \<C>\<close> \<open>v \<in> T\<close>] by simp
+    also have vS: "setdist {v} S \<le> dist a v"
+      by (simp add: setdist_le_dist setdist_sym \<open>a \<in> S\<close>)
+    finally have VTV: "dist (\<V> T) v < dist a v" .
+    have VTS: "setdist {\<V> T} S \<le> 2 * dist a v"
+      using sdle that vS by force
+    have "dist a (\<A> T) \<le> dist a v + dist v (\<V> T) + dist (\<V> T) (\<A> T)"
+      by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le)
+    also have "... \<le> dist a v + dist a v + dist (\<V> T) (\<A> T)"
+      using VTV by (simp add: dist_commute)
+    also have "... \<le> 2 * dist a v + 2 * setdist {\<V> T} S"
+      using VA [OF \<open>T \<in> \<C>\<close>] by auto
+    finally show ?thesis
+      using VTS by linarith
+  qed
+  obtain H :: "['a set, 'a] \<Rightarrow> real"
+    where Hcont: "\<And>Z. Z \<in> \<C> \<Longrightarrow> continuous_on (U-S) (H Z)"
+      and Hge0: "\<And>Z x. \<lbrakk>Z \<in> \<C>; x \<in> U-S\<rbrakk> \<Longrightarrow> 0 \<le> H Z x"
+      and Heq0: "\<And>x Z. \<lbrakk>Z \<in> \<C>; x \<in> U-S; x \<notin> Z\<rbrakk> \<Longrightarrow> H Z x = 0"
+      and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_setsum (\<lambda>W. H W x) \<C> = 1"
+      and Hfin: "\<And>x. x \<in> U-S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. H U x \<noteq> 0}"
+    apply (rule subordinate_partition_of_unity [OF USsub _ fin])
+    using nbrhd by auto
+  define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_setsum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>"
+  show ?thesis
+  proof (rule that)
+    show "continuous_on U g"
+    proof (clarsimp simp: continuous_on_eq_continuous_within)
+      fix a assume "a \<in> U"
+      show "continuous (at a within U) g"
+      proof (cases "a \<in> S")
+        case True show ?thesis
+        proof (clarsimp simp add: continuous_within_topological)
+          fix W
+          assume "open W" "g a \<in> W"
+          then obtain e where "0 < e" and e: "ball (f a) e \<subseteq> W"
+            using openE True g_def by auto
+          have "continuous (at a within S) f"
+            using True contf continuous_on_eq_continuous_within by blast
+          then obtain d where "0 < d"
+                        and d: "\<And>x. \<lbrakk>x \<in> S; dist x a < d\<rbrakk> \<Longrightarrow> dist (f x) (f a) < e"
+            using continuous_within_eps_delta \<open>0 < e\<close> by force
+          have "g y \<in> ball (f a) e" if "y \<in> U" and y: "y \<in> ball a (d / 6)" for y
+          proof (cases "y \<in> S")
+            case True
+            then have "dist (f a) (f y) < e"
+              by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d)
+            then show ?thesis
+              by (simp add: True g_def)
+          next
+            case False
+            have *: "dist (f (\<A> T)) (f a) < e" if "T \<in> \<C>" "H T y \<noteq> 0" for T
+            proof -
+              have "y \<in> T"
+                using Heq0 that False \<open>y \<in> U\<close> by blast
+              have "dist (\<A> T) a < d"
+                using d6 [OF \<open>T \<in> \<C>\<close> \<open>y \<in> T\<close> \<open>a \<in> S\<close>] y
+                by (simp add: dist_commute mult.commute)
+              then show ?thesis
+                using VA [OF \<open>T \<in> \<C>\<close>] by (auto simp: d)
+            qed
+            have "supp_setsum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e"
+              apply (rule convex_supp_setsum [OF convex_ball])
+              apply (simp_all add: False H1 Hge0 \<open>y \<in> U\<close>)
+              by (metis dist_commute *)
+            then show ?thesis
+              by (simp add: False g_def)
+          qed
+          then show "\<exists>A. open A \<and> a \<in> A \<and> (\<forall>y\<in>U. y \<in> A \<longrightarrow> g y \<in> W)"
+            apply (rule_tac x = "ball a (d / 6)" in exI)
+            using e \<open>0 < d\<close> by fastforce
+        qed
+      next
+        case False
+        obtain N where N: "open N" "a \<in> N"
+                   and finN: "finite {U \<in> \<C>. \<exists>a\<in>N. H U a \<noteq> 0}"
+          using Hfin False \<open>a \<in> U\<close> by auto
+        have oUS: "openin (subtopology euclidean U) (U - S)"
+          using cloin by (simp add: openin_diff)
+        have HcontU: "continuous (at a within U) (H T)" if "T \<in> \<C>" for T
+          using Hcont [OF \<open>T \<in> \<C>\<close>] False  \<open>a \<in> U\<close> \<open>T \<in> \<C>\<close>
+          apply (simp add: continuous_on_eq_continuous_within continuous_within)
+          apply (rule Lim_transform_within_set)
+          using oUS
+            apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+
+          done
+        show ?thesis
+        proof (rule continuous_transform_within_openin [OF _ oUS])
+          show "continuous (at a within U) (\<lambda>x. supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)"
+          proof (rule continuous_transform_within_openin)
+            show "continuous (at a within U)
+                    (\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))"
+              by (force intro: continuous_intros HcontU)+
+          next
+            show "openin (subtopology euclidean U) ((U - S) \<inter> N)"
+              using N oUS openin_trans by blast
+          next
+            show "a \<in> (U - S) \<inter> N" using False \<open>a \<in> U\<close> N by blast
+          next
+            show "\<And>x. x \<in> (U - S) \<inter> N \<Longrightarrow>
+                         (\<Sum>T \<in> {U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))
+                         = supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>"
+              by (auto simp: supp_setsum_def support_on_def
+                       intro: setsum.mono_neutral_right [OF finN])
+          qed
+        next
+          show "a \<in> U - S" using False \<open>a \<in> U\<close> by blast
+        next
+          show "\<And>x. x \<in> U - S \<Longrightarrow> supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x"
+            by (simp add: g_def)
+        qed
+      qed
+    qed
+    show "g ` U \<subseteq> C"
+      using \<open>f ` S \<subseteq> C\<close> VA
+      by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \<open>convex C\<close>] H1)
+    show "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+      by (simp add: g_def)
+  qed
+qed
+
+
+corollary Tietze:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+      and "0 \<le> B"
+      and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> B"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+                  "\<And>x. x \<in> U \<Longrightarrow> norm(g x) \<le> B"
+using assms
+by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f])
+
+corollary Tietze_closed_interval:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+      and "cbox a b \<noteq> {}"
+      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
+apply (rule Dugundji [of "cbox a b" U S f])
+using assms by auto
+
+corollary Tietze_closed_interval_1:
+  fixes f :: "'a::euclidean_space \<Rightarrow> real"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+      and "a \<le> b"
+      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
+apply (rule Dugundji [of "cbox a b" U S f])
+using assms by (auto simp: image_subset_iff)
+
+corollary Tietze_open_interval:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+      and "box a b \<noteq> {}"
+      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
+apply (rule Dugundji [of "box a b" U S f])
+using assms by auto
+
+corollary Tietze_open_interval_1:
+  fixes f :: "'a::euclidean_space \<Rightarrow> real"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+      and "a < b"
+      and no: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
+apply (rule Dugundji [of "box a b" U S f])
+using assms by (auto simp: image_subset_iff)
+
+corollary Tietze_unbounded:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
+  assumes "continuous_on S f"
+      and "closedin (subtopology euclidean U) S"
+  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
+apply (rule Dugundji [of UNIV U S f])
+using assms by auto
+
+end
--- a/src/HOL/Multivariate_Analysis/Extension.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,547 +0,0 @@
-(*  Title:      HOL/Multivariate_Analysis/Extension.thy
-    Authors:    LC Paulson, based on material from HOL Light
-*)
-
-section \<open>Continuous extensions of functions: Urysohn's lemma, Dugundji extension theorem, Tietze\<close>
-
-theory Extension
-imports Convex_Euclidean_Space
-begin
-
-subsection\<open>Partitions of unity subordinate to locally finite open coverings\<close>
-
-text\<open>A difference from HOL Light: all summations over infinite sets equal zero,
-   so the "support" must be made explicit in the summation below!\<close>
-
-proposition subordinate_partition_of_unity:
-  fixes S :: "'a :: euclidean_space set"
-  assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
-      and fin: "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
-  obtains F :: "['a set, 'a] \<Rightarrow> real"
-    where "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x \<in> S. 0 \<le> F U x)"
-      and "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
-      and "\<And>x. x \<in> S \<Longrightarrow> supp_setsum (\<lambda>W. F W x) \<C> = 1"
-      and "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}"
-proof (cases "\<exists>W. W \<in> \<C> \<and> S \<subseteq> W")
-  case True
-    then obtain W where "W \<in> \<C>" "S \<subseteq> W" by metis
-    then show ?thesis
-      apply (rule_tac F = "\<lambda>V x. if V = W then 1 else 0" in that)
-      apply (auto simp: continuous_on_const supp_setsum_def support_on_def)
-      done
-next
-  case False
-    have nonneg: "0 \<le> supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" for x
-      by (simp add: supp_setsum_def setsum_nonneg)
-    have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> V" for V x
-    proof -
-      have "closedin (subtopology euclidean S) (S - V)"
-        by (simp add: Diff_Diff_Int Diff_subset closedin_def opC openin_open_Int \<open>V \<in> \<C>\<close>)
-      with that False setdist_eq_0_closedin [of S "S-V" x] setdist_pos_le [of "{x}" "S - V"]
-        show ?thesis
-          by (simp add: order_class.order.order_iff_strict)
-    qed
-    have ss_pos: "0 < supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x
-    proof -
-      obtain U where "U \<in> \<C>" "x \<in> U" using \<open>x \<in> S\<close> \<open>S \<subseteq> \<Union>\<C>\<close>
-        by blast
-      obtain V where "open V" "x \<in> V" "finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
-        using \<open>x \<in> S\<close> fin by blast
-      then have *: "finite {A \<in> \<C>. \<not> S \<subseteq> A \<and> x \<notin> closure (S - A)}"
-        using closure_def that by (blast intro: rev_finite_subset)
-      have "x \<notin> closure (S - U)"
-        by (metis \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> less_irrefl sd_pos setdist_eq_0_sing_1 that)
-      then show ?thesis
-        apply (simp add: setdist_eq_0_sing_1 supp_setsum_def support_on_def)
-        apply (rule ordered_comm_monoid_add_class.setsum_pos2 [OF *, of U])
-        using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> False
-        apply (auto simp: setdist_pos_le sd_pos that)
-        done
-    qed
-    define F where
-      "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>
-                 else 0"
-    show ?thesis
-    proof (rule_tac F = F in that)
-      have "continuous_on S (F U)" if "U \<in> \<C>" for U
-      proof -
-        have *: "continuous_on S (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
-        proof (clarsimp simp add: continuous_on_eq_continuous_within)
-          fix x assume "x \<in> S"
-          then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}"
-            using assms by blast
-          then have OSX: "openin (subtopology euclidean S) (S \<inter> X)" by blast
-          have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow>
-                     (\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V))
-                     = supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>"
-            apply (simp add: supp_setsum_def)
-            apply (rule setsum.mono_neutral_right [OF finX])
-            apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff)
-            apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE)
-            done
-          show "continuous (at x within S) (\<lambda>x. supp_setsum (\<lambda>V. setdist {x} (S - V)) \<C>)"
-            apply (rule continuous_transform_within_openin
-                     [where f = "\<lambda>x. (setsum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})"
-                        and S ="S \<inter> X"])
-            apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+
-            apply (simp add: sumeq)
-            done
-        qed
-        show ?thesis
-          apply (simp add: F_def)
-          apply (rule continuous_intros *)+
-          using ss_pos apply force
-          done
-      qed
-      moreover have "\<lbrakk>U \<in> \<C>; x \<in> S\<rbrakk> \<Longrightarrow> 0 \<le> F U x" for U x
-        using nonneg [of x] by (simp add: F_def divide_simps setdist_pos_le)
-      ultimately show "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x\<in>S. 0 \<le> F U x)"
-        by metis
-    next
-      show "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
-        by (simp add: setdist_eq_0_sing_1 closure_def F_def)
-    next
-      show "supp_setsum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x
-        using that ss_pos [OF that]
-        by (simp add: F_def divide_simps supp_setsum_divide_distrib [symmetric])
-    next
-      show "\<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" if "x \<in> S" for x
-        using fin [OF that] that
-        by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset)
-    qed
-qed
-
-
-subsection\<open>Urysohn's lemma (for Euclidean spaces, where the proof is easy using distances)\<close>
-
-lemma Urysohn_both_ne:
-  assumes US: "closedin (subtopology euclidean U) S"
-      and UT: "closedin (subtopology euclidean U) T"
-      and "S \<inter> T = {}" "S \<noteq> {}" "T \<noteq> {}" "a \<noteq> b"
-  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-    where "continuous_on U f"
-          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
-          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
-          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
-proof -
-  have S0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} S = 0 \<longleftrightarrow> x \<in> S"
-    using \<open>S \<noteq> {}\<close>  US setdist_eq_0_closedin  by auto
-  have T0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} T = 0 \<longleftrightarrow> x \<in> T"
-    using \<open>T \<noteq> {}\<close>  UT setdist_eq_0_closedin  by auto
-  have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \<in> U" for x
-  proof -
-    have "~ (setdist {x} S = 0 \<and> setdist {x} T = 0)"
-      using assms by (metis IntI empty_iff setdist_eq_0_closedin that)
-    then show ?thesis
-      by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le)
-  qed
-  define f where "f \<equiv> \<lambda>x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)"
-  show ?thesis
-  proof (rule_tac f = f in that)
-    show "continuous_on U f"
-      using sdpos unfolding f_def
-      by (intro continuous_intros | force)+
-    show "f x \<in> closed_segment a b" if "x \<in> U" for x
-        unfolding f_def
-      apply (simp add: closed_segment_def)
-      apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI)
-      using sdpos that apply (simp add: algebra_simps)
-      done
-    show "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
-      using S0 \<open>a \<noteq> b\<close> f_def sdpos by force
-    show "(f x = b \<longleftrightarrow> x \<in> T)" if "x \<in> U" for x
-    proof -
-      have "f x = b \<longleftrightarrow> (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1"
-        unfolding f_def
-        apply (rule iffI)
-         apply (metis  \<open>a \<noteq> b\<close> add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force)
-        done
-      also have "...  \<longleftrightarrow> setdist {x} T = 0 \<and> setdist {x} S \<noteq> 0"
-        using sdpos that
-        by (simp add: divide_simps) linarith
-      also have "...  \<longleftrightarrow> x \<in> T"
-        using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>S \<inter> T = {}\<close> that
-        by (force simp: S0 T0)
-      finally show ?thesis .
-    qed
-  qed
-qed
-
-proposition Urysohn_local_strong:
-  assumes US: "closedin (subtopology euclidean U) S"
-      and UT: "closedin (subtopology euclidean U) T"
-      and "S \<inter> T = {}" "a \<noteq> b"
-  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    where "continuous_on U f"
-          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
-          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
-          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
-proof (cases "S = {}")
-  case True show ?thesis
-  proof (cases "T = {}")
-    case True show ?thesis
-    proof (rule_tac f = "\<lambda>x. midpoint a b" in that)
-      show "continuous_on U (\<lambda>x. midpoint a b)"
-        by (intro continuous_intros)
-      show "midpoint a b \<in> closed_segment a b"
-        using csegment_midpoint_subset by blast
-      show "(midpoint a b = a) = (x \<in> S)" for x
-        using \<open>S = {}\<close> \<open>a \<noteq> b\<close> by simp
-      show "(midpoint a b = b) = (x \<in> T)" for x
-        using \<open>T = {}\<close> \<open>a \<noteq> b\<close> by simp
-    qed
-  next
-    case False
-    show ?thesis
-    proof (cases "T = U")
-      case True with \<open>S = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
-        by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const)
-    next
-      case False
-      with UT closedin_subset obtain c where c: "c \<in> U" "c \<notin> T"
-        by fastforce
-      obtain f where f: "continuous_on U f"
-                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment (midpoint a b) b"
-                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
-                "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
-        apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"])
-        using c \<open>T \<noteq> {}\<close> assms apply simp_all
-        done
-      show ?thesis
-        apply (rule_tac f=f in that)
-        using \<open>S = {}\<close> \<open>T \<noteq> {}\<close> f csegment_midpoint_subset notin_segment_midpoint [OF \<open>a \<noteq> b\<close>]
-        apply force+
-        done
-    qed
-  qed
-next
-  case False
-  show ?thesis
-  proof (cases "T = {}")
-    case True show ?thesis
-    proof (cases "S = U")
-      case True with \<open>T = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
-        by (rule_tac f = "\<lambda>x. a" in that) (auto simp: continuous_on_const)
-    next
-      case False
-      with US closedin_subset obtain c where c: "c \<in> U" "c \<notin> S"
-        by fastforce
-      obtain f where f: "continuous_on U f"
-                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a (midpoint a b)"
-                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
-                "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
-        apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"])
-        using c \<open>S \<noteq> {}\<close> assms apply simp_all
-        apply (metis midpoint_eq_endpoint)
-        done
-      show ?thesis
-        apply (rule_tac f=f in that)
-        using \<open>S \<noteq> {}\<close> \<open>T = {}\<close> f  \<open>a \<noteq> b\<close>
-        apply simp_all
-        apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff)
-        apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint)
-        done
-    qed
-  next
-    case False
-    show ?thesis
-      using Urysohn_both_ne [OF US UT \<open>S \<inter> T = {}\<close> \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>a \<noteq> b\<close>] that
-      by blast
-  qed
-qed
-
-lemma Urysohn_local:
-  assumes US: "closedin (subtopology euclidean U) S"
-      and UT: "closedin (subtopology euclidean U) T"
-      and "S \<inter> T = {}"
-  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    where "continuous_on U f"
-          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
-          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
-          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
-proof (cases "a = b")
-  case True then show ?thesis
-    by (rule_tac f = "\<lambda>x. b" in that) (auto simp: continuous_on_const)
-next
-  case False
-  then show ?thesis
-    apply (rule Urysohn_local_strong [OF assms])
-    apply (erule that, assumption)
-    apply (meson US closedin_singleton closedin_trans)
-    apply (meson UT closedin_singleton closedin_trans)
-    done
-qed
-
-lemma Urysohn_strong:
-  assumes US: "closed S"
-      and UT: "closed T"
-      and "S \<inter> T = {}" "a \<noteq> b"
-  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    where "continuous_on UNIV f"
-          "\<And>x. f x \<in> closed_segment a b"
-          "\<And>x. f x = a \<longleftrightarrow> x \<in> S"
-          "\<And>x. f x = b \<longleftrightarrow> x \<in> T"
-apply (rule Urysohn_local_strong [of UNIV S T])
-using assms
-apply (auto simp: closed_closedin)
-done
-
-proposition Urysohn:
-  assumes US: "closed S"
-      and UT: "closed T"
-      and "S \<inter> T = {}"
-  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-    where "continuous_on UNIV f"
-          "\<And>x. f x \<in> closed_segment a b"
-          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
-          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
-apply (rule Urysohn_local [of UNIV S T a b])
-using assms
-apply (auto simp: closed_closedin)
-done
-
-
-subsection\<open> The Dugundji extension theorem, and Tietze variants as corollaries.\<close>
-
-text\<open>J. Dugundji. An extension of Tietze's theorem. Pacific J. Math. Volume 1, Number 3 (1951), 353-367.
-http://projecteuclid.org/euclid.pjm/1103052106\<close>
-
-theorem Dugundji:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
-  assumes "convex C" "C \<noteq> {}"
-      and cloin: "closedin (subtopology euclidean U) S"
-      and contf: "continuous_on S f" and "f ` S \<subseteq> C"
-  obtains g where "continuous_on U g" "g ` U \<subseteq> C"
-                  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-proof (cases "S = {}")
-  case True then show thesis
-    apply (rule_tac g="\<lambda>x. @y. y \<in> C" in that)
-      apply (rule continuous_intros)
-     apply (meson all_not_in_conv \<open>C \<noteq> {}\<close> image_subsetI someI_ex, simp)
-    done
-next
-  case False
-  then have sd_pos: "\<And>x. \<lbrakk>x \<in> U; x \<notin> S\<rbrakk> \<Longrightarrow> 0 < setdist {x} S"
-    using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce
-  define \<B> where "\<B> = {ball x (setdist {x} S / 2) |x. x \<in> U - S}"
-  have [simp]: "\<And>T. T \<in> \<B> \<Longrightarrow> open T"
-    by (auto simp: \<B>_def)
-  have USS: "U - S \<subseteq> \<Union>\<B>"
-    by (auto simp: sd_pos \<B>_def)
-  obtain \<C> where USsub: "U - S \<subseteq> \<Union>\<C>"
-       and nbrhd: "\<And>U. U \<in> \<C> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<B> \<and> U \<subseteq> T)"
-       and fin: "\<And>x. x \<in> U - S
-                     \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> \<C> \<and> U \<inter> V \<noteq> {}}"
-    using paracompact [OF USS] by auto
-  have "\<exists>v a. v \<in> U \<and> v \<notin> S \<and> a \<in> S \<and>
-              T \<subseteq> ball v (setdist {v} S / 2) \<and>
-              dist v a \<le> 2 * setdist {v} S" if "T \<in> \<C>" for T
-  proof -
-    obtain v where v: "T \<subseteq> ball v (setdist {v} S / 2)" "v \<in> U" "v \<notin> S"
-      using \<open>T \<in> \<C>\<close> nbrhd by (force simp: \<B>_def)
-    then obtain a where "a \<in> S" "dist v a < 2 * setdist {v} S"
-      using setdist_ltE [of "{v}" S "2 * setdist {v} S"]
-      using False sd_pos by force
-    with v show ?thesis
-      apply (rule_tac x=v in exI)
-      apply (rule_tac x=a in exI, auto)
-      done
-  qed
-  then obtain \<V> \<A> where
-    VA: "\<And>T. T \<in> \<C> \<Longrightarrow> \<V> T \<in> U \<and> \<V> T \<notin> S \<and> \<A> T \<in> S \<and>
-              T \<subseteq> ball (\<V> T) (setdist {\<V> T} S / 2) \<and>
-              dist (\<V> T) (\<A> T) \<le> 2 * setdist {\<V> T} S"
-    by metis
-  have sdle: "setdist {\<V> T} S \<le> 2 * setdist {v} S" if "T \<in> \<C>" "v \<in> T" for T v
-    using setdist_Lipschitz [of "\<V> T" S v] VA [OF \<open>T \<in> \<C>\<close>] \<open>v \<in> T\<close> by auto
-  have d6: "dist a (\<A> T) \<le> 6 * dist a v" if "T \<in> \<C>" "v \<in> T" "a \<in> S" for T v a
-  proof -
-    have "dist (\<V> T) v < setdist {\<V> T} S / 2"
-      using that VA mem_ball by blast
-    also have "... \<le> setdist {v} S"
-      using sdle [OF \<open>T \<in> \<C>\<close> \<open>v \<in> T\<close>] by simp
-    also have vS: "setdist {v} S \<le> dist a v"
-      by (simp add: setdist_le_dist setdist_sym \<open>a \<in> S\<close>)
-    finally have VTV: "dist (\<V> T) v < dist a v" .
-    have VTS: "setdist {\<V> T} S \<le> 2 * dist a v"
-      using sdle that vS by force
-    have "dist a (\<A> T) \<le> dist a v + dist v (\<V> T) + dist (\<V> T) (\<A> T)"
-      by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le)
-    also have "... \<le> dist a v + dist a v + dist (\<V> T) (\<A> T)"
-      using VTV by (simp add: dist_commute)
-    also have "... \<le> 2 * dist a v + 2 * setdist {\<V> T} S"
-      using VA [OF \<open>T \<in> \<C>\<close>] by auto
-    finally show ?thesis
-      using VTS by linarith
-  qed
-  obtain H :: "['a set, 'a] \<Rightarrow> real"
-    where Hcont: "\<And>Z. Z \<in> \<C> \<Longrightarrow> continuous_on (U-S) (H Z)"
-      and Hge0: "\<And>Z x. \<lbrakk>Z \<in> \<C>; x \<in> U-S\<rbrakk> \<Longrightarrow> 0 \<le> H Z x"
-      and Heq0: "\<And>x Z. \<lbrakk>Z \<in> \<C>; x \<in> U-S; x \<notin> Z\<rbrakk> \<Longrightarrow> H Z x = 0"
-      and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_setsum (\<lambda>W. H W x) \<C> = 1"
-      and Hfin: "\<And>x. x \<in> U-S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. H U x \<noteq> 0}"
-    apply (rule subordinate_partition_of_unity [OF USsub _ fin])
-    using nbrhd by auto
-  define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_setsum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>"
-  show ?thesis
-  proof (rule that)
-    show "continuous_on U g"
-    proof (clarsimp simp: continuous_on_eq_continuous_within)
-      fix a assume "a \<in> U"
-      show "continuous (at a within U) g"
-      proof (cases "a \<in> S")
-        case True show ?thesis
-        proof (clarsimp simp add: continuous_within_topological)
-          fix W
-          assume "open W" "g a \<in> W"
-          then obtain e where "0 < e" and e: "ball (f a) e \<subseteq> W"
-            using openE True g_def by auto
-          have "continuous (at a within S) f"
-            using True contf continuous_on_eq_continuous_within by blast
-          then obtain d where "0 < d"
-                        and d: "\<And>x. \<lbrakk>x \<in> S; dist x a < d\<rbrakk> \<Longrightarrow> dist (f x) (f a) < e"
-            using continuous_within_eps_delta \<open>0 < e\<close> by force
-          have "g y \<in> ball (f a) e" if "y \<in> U" and y: "y \<in> ball a (d / 6)" for y
-          proof (cases "y \<in> S")
-            case True
-            then have "dist (f a) (f y) < e"
-              by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d)
-            then show ?thesis
-              by (simp add: True g_def)
-          next
-            case False
-            have *: "dist (f (\<A> T)) (f a) < e" if "T \<in> \<C>" "H T y \<noteq> 0" for T
-            proof -
-              have "y \<in> T"
-                using Heq0 that False \<open>y \<in> U\<close> by blast
-              have "dist (\<A> T) a < d"
-                using d6 [OF \<open>T \<in> \<C>\<close> \<open>y \<in> T\<close> \<open>a \<in> S\<close>] y
-                by (simp add: dist_commute mult.commute)
-              then show ?thesis
-                using VA [OF \<open>T \<in> \<C>\<close>] by (auto simp: d)
-            qed
-            have "supp_setsum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e"
-              apply (rule convex_supp_setsum [OF convex_ball])
-              apply (simp_all add: False H1 Hge0 \<open>y \<in> U\<close>)
-              by (metis dist_commute *)
-            then show ?thesis
-              by (simp add: False g_def)
-          qed
-          then show "\<exists>A. open A \<and> a \<in> A \<and> (\<forall>y\<in>U. y \<in> A \<longrightarrow> g y \<in> W)"
-            apply (rule_tac x = "ball a (d / 6)" in exI)
-            using e \<open>0 < d\<close> by fastforce
-        qed
-      next
-        case False
-        obtain N where N: "open N" "a \<in> N"
-                   and finN: "finite {U \<in> \<C>. \<exists>a\<in>N. H U a \<noteq> 0}"
-          using Hfin False \<open>a \<in> U\<close> by auto
-        have oUS: "openin (subtopology euclidean U) (U - S)"
-          using cloin by (simp add: openin_diff)
-        have HcontU: "continuous (at a within U) (H T)" if "T \<in> \<C>" for T
-          using Hcont [OF \<open>T \<in> \<C>\<close>] False  \<open>a \<in> U\<close> \<open>T \<in> \<C>\<close>
-          apply (simp add: continuous_on_eq_continuous_within continuous_within)
-          apply (rule Lim_transform_within_set)
-          using oUS
-            apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+
-          done
-        show ?thesis
-        proof (rule continuous_transform_within_openin [OF _ oUS])
-          show "continuous (at a within U) (\<lambda>x. supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)"
-          proof (rule continuous_transform_within_openin)
-            show "continuous (at a within U)
-                    (\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))"
-              by (force intro: continuous_intros HcontU)+
-          next
-            show "openin (subtopology euclidean U) ((U - S) \<inter> N)"
-              using N oUS openin_trans by blast
-          next
-            show "a \<in> (U - S) \<inter> N" using False \<open>a \<in> U\<close> N by blast
-          next
-            show "\<And>x. x \<in> (U - S) \<inter> N \<Longrightarrow>
-                         (\<Sum>T \<in> {U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))
-                         = supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>"
-              by (auto simp: supp_setsum_def support_on_def
-                       intro: setsum.mono_neutral_right [OF finN])
-          qed
-        next
-          show "a \<in> U - S" using False \<open>a \<in> U\<close> by blast
-        next
-          show "\<And>x. x \<in> U - S \<Longrightarrow> supp_setsum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x"
-            by (simp add: g_def)
-        qed
-      qed
-    qed
-    show "g ` U \<subseteq> C"
-      using \<open>f ` S \<subseteq> C\<close> VA
-      by (fastforce simp: g_def Hge0 intro!: convex_supp_setsum [OF \<open>convex C\<close>] H1)
-    show "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-      by (simp add: g_def)
-  qed
-qed
-
-
-corollary Tietze:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-      and "0 \<le> B"
-      and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> B"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-                  "\<And>x. x \<in> U \<Longrightarrow> norm(g x) \<le> B"
-using assms
-by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f])
-
-corollary Tietze_closed_interval:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-      and "cbox a b \<noteq> {}"
-      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
-apply (rule Dugundji [of "cbox a b" U S f])
-using assms by auto
-
-corollary Tietze_closed_interval_1:
-  fixes f :: "'a::euclidean_space \<Rightarrow> real"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-      and "a \<le> b"
-      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
-apply (rule Dugundji [of "cbox a b" U S f])
-using assms by (auto simp: image_subset_iff)
-
-corollary Tietze_open_interval:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-      and "box a b \<noteq> {}"
-      and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
-apply (rule Dugundji [of "box a b" U S f])
-using assms by auto
-
-corollary Tietze_open_interval_1:
-  fixes f :: "'a::euclidean_space \<Rightarrow> real"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-      and "a < b"
-      and no: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-                  "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
-apply (rule Dugundji [of "box a b" U S f])
-using assms by (auto simp: image_subset_iff)
-
-corollary Tietze_unbounded:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_inner"
-  assumes "continuous_on S f"
-      and "closedin (subtopology euclidean U) S"
-  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
-apply (rule Dugundji [of UNIV U S f])
-using assms by auto
-
-end
--- a/src/HOL/Multivariate_Analysis/Fashoda.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1136 +0,0 @@
-(*  Author:     John Harrison
-    Author:     Robert Himmelmann, TU Muenchen (translation from HOL light)
-*)
-
-section \<open>Fashoda meet theorem\<close>
-
-theory Fashoda
-imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space
-begin
-
-subsection \<open>Bijections between intervals.\<close>
-
-definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
-  where "interval_bij =
-    (\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))"
-
-lemma interval_bij_affine:
-  "interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) +
-    (\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))"
-  by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
-    field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong)
-
-lemma continuous_interval_bij:
-  fixes a b :: "'a::euclidean_space"
-  shows "continuous (at x) (interval_bij (a, b) (u, v))"
-  by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros)
-
-lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))"
-  apply(rule continuous_at_imp_continuous_on)
-  apply (rule, rule continuous_interval_bij)
-  done
-
-lemma in_interval_interval_bij:
-  fixes a b u v x :: "'a::euclidean_space"
-  assumes "x \<in> cbox a b"
-    and "cbox u v \<noteq> {}"
-  shows "interval_bij (a, b) (u, v) x \<in> cbox u v"
-  apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong)
-  apply safe
-proof -
-  fix i :: 'a
-  assume i: "i \<in> Basis"
-  have "cbox a b \<noteq> {}"
-    using assms by auto
-  with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i"
-    using assms(2) by (auto simp add: box_eq_empty)
-  have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i"
-    using assms(1)[unfolded mem_box] using i by auto
-  have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
-    using * x by auto
-  then show "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
-    using * by auto
-  have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)"
-    apply (rule mult_right_mono)
-    unfolding divide_le_eq_1
-    using * x
-    apply auto
-    done
-  then show "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i"
-    using * by auto
-qed
-
-lemma interval_bij_bij:
-  "\<forall>(i::'a::euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow>
-    interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x"
-  by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a])
-
-lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i" 
-  shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
-  using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis)
-
-
-subsection \<open>Fashoda meet theorem\<close>
-
-lemma infnorm_2:
-  fixes x :: "real^2"
-  shows "infnorm x = max \<bar>x$1\<bar> \<bar>x$2\<bar>"
-  unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto
-
-lemma infnorm_eq_1_2:
-  fixes x :: "real^2"
-  shows "infnorm x = 1 \<longleftrightarrow>
-    \<bar>x$1\<bar> \<le> 1 \<and> \<bar>x$2\<bar> \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1)"
-  unfolding infnorm_2 by auto
-
-lemma infnorm_eq_1_imp:
-  fixes x :: "real^2"
-  assumes "infnorm x = 1"
-  shows "\<bar>x$1\<bar> \<le> 1" and "\<bar>x$2\<bar> \<le> 1"
-  using assms unfolding infnorm_eq_1_2 by auto
-
-lemma fashoda_unit:
-  fixes f g :: "real \<Rightarrow> real^2"
-  assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1"
-    and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1"
-    and "continuous_on {-1 .. 1} f"
-    and "continuous_on {-1 .. 1} g"
-    and "f (- 1)$1 = - 1"
-    and "f 1$1 = 1" "g (- 1) $2 = -1"
-    and "g 1 $2 = 1"
-  shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t"
-proof (rule ccontr)
-  assume "\<not> ?thesis"
-  note as = this[unfolded bex_simps,rule_format]
-  define sqprojection
-    where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2"
-  define negatex :: "real^2 \<Rightarrow> real^2"
-    where "negatex x = (vector [-(x$1), x$2])" for x
-  have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z"
-    unfolding negatex_def infnorm_2 vector_2 by auto
-  have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1"
-    unfolding sqprojection_def
-    unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR]
-    unfolding abs_inverse real_abs_infnorm
-    apply (subst infnorm_eq_0[symmetric])
-    apply auto
-    done
-  let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w"
-  have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}"
-    apply (rule set_eqI)
-    unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart
-    apply rule
-    defer
-    apply (rule_tac x="vec x" in exI)
-    apply auto
-    done
-  {
-    fix x
-    assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` (cbox (- 1) (1::real^2))"
-    then obtain w :: "real^2" where w:
-        "w \<in> cbox (- 1) 1"
-        "x = (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w"
-      unfolding image_iff ..
-    then have "x \<noteq> 0"
-      using as[of "w$1" "w$2"]
-      unfolding mem_interval_cart atLeastAtMost_iff
-      by auto
-  } note x0 = this
-  have 21: "\<And>i::2. i \<noteq> 1 \<Longrightarrow> i = 2"
-    using UNIV_2 by auto
-  have 1: "box (- 1) (1::real^2) \<noteq> {}"
-    unfolding interval_eq_empty_cart by auto
-  have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)"
-    apply (intro continuous_intros continuous_on_component)
-    unfolding *
-    apply (rule assms)+
-    apply (subst sqprojection_def)
-    apply (intro continuous_intros)
-    apply (simp add: infnorm_eq_0 x0)
-    apply (rule linear_continuous_on)
-  proof -
-    show "bounded_linear negatex"
-      apply (rule bounded_linearI')
-      unfolding vec_eq_iff
-    proof (rule_tac[!] allI)
-      fix i :: 2
-      fix x y :: "real^2"
-      fix c :: real
-      show "negatex (x + y) $ i =
-        (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i"
-        apply -
-        apply (case_tac[!] "i\<noteq>1")
-        prefer 3
-        apply (drule_tac[1-2] 21) 
-        unfolding negatex_def
-        apply (auto simp add:vector_2)
-        done
-    qed
-  qed
-  have 3: "(negatex \<circ> sqprojection \<circ> ?F) ` cbox (-1) 1 \<subseteq> cbox (-1) 1"
-    unfolding subset_eq
-  proof (rule, goal_cases)
-    case (1 x)
-    then obtain y :: "real^2" where y:
-        "y \<in> cbox (- 1) 1"
-        "x = (negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) y"
-      unfolding image_iff ..
-    have "?F y \<noteq> 0"
-      apply (rule x0)
-      using y(1)
-      apply auto
-      done
-    then have *: "infnorm (sqprojection (?F y)) = 1"
-      unfolding y o_def
-      by - (rule lem2[rule_format])
-    have "infnorm x = 1"
-      unfolding *[symmetric] y o_def
-      by (rule lem1[rule_format])
-    then show "x \<in> cbox (-1) 1"
-      unfolding mem_interval_cart interval_cbox_cart infnorm_2
-      apply -
-      apply rule
-    proof -
-      fix i
-      assume "max \<bar>x $ 1\<bar> \<bar>x $ 2\<bar> = 1"
-      then show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i"
-        apply (cases "i = 1")
-        defer
-        apply (drule 21)
-        apply auto
-        done
-    qed
-  qed
-  obtain x :: "real^2" where x:
-      "x \<in> cbox (- 1) 1"
-      "(negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) x = x"
-    apply (rule brouwer_weak[of "cbox (- 1) (1::real^2)" "negatex \<circ> sqprojection \<circ> ?F"])
-    apply (rule compact_cbox convex_box)+
-    unfolding interior_cbox
-    apply (rule 1 2 3)+
-    apply blast
-    done
-  have "?F x \<noteq> 0"
-    apply (rule x0)
-    using x(1)
-    apply auto
-    done
-  then have *: "infnorm (sqprojection (?F x)) = 1"
-    unfolding o_def
-    by (rule lem2[rule_format])
-  have nx: "infnorm x = 1"
-    apply (subst x(2)[symmetric])
-    unfolding *[symmetric] o_def
-    apply (rule lem1[rule_format])
-    done
-  have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)"
-    and "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)"
-    apply -
-    apply (rule_tac[!] allI impI)+
-  proof -
-    fix x :: "real^2"
-    fix i :: 2
-    assume x: "x \<noteq> 0"
-    have "inverse (infnorm x) > 0"
-      using x[unfolded infnorm_pos_lt[symmetric]] by auto
-    then show "(0 < sqprojection x $ i) = (0 < x $ i)"
-      and "(sqprojection x $ i < 0) = (x $ i < 0)"
-      unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def
-      unfolding zero_less_mult_iff mult_less_0_iff
-      by (auto simp add: field_simps)
-  qed
-  note lem3 = this[rule_format]
-  have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}"
-    using x(1) unfolding mem_interval_cart by auto
-  then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0"
-    unfolding right_minus_eq
-    apply -
-    apply (rule as)
-    apply auto
-    done
-  have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1"
-    using nx unfolding infnorm_eq_1_2 by auto 
-  then show False
-  proof -
-    fix P Q R S 
-    presume "P \<or> Q \<or> R \<or> S"
-      and "P \<Longrightarrow> False"
-      and "Q \<Longrightarrow> False"
-      and "R \<Longrightarrow> False"
-      and "S \<Longrightarrow> False"
-    then show False by auto
-  next
-    assume as: "x$1 = 1"
-    then have *: "f (x $ 1) $ 1 = 1"
-      using assms(6) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "g (x $ 2) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(2)[unfolded subset_eq,rule_format])
-      apply auto
-      done
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
-      apply (erule_tac x=1 in allE)
-      apply auto
-      done
-  next
-    assume as: "x$1 = -1"
-    then have *: "f (x $ 1) $ 1 = - 1"
-      using assms(5) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "g (x $ 2) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(2)[unfolded subset_eq,rule_format])
-      apply auto
-      done
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart 
-      apply (erule_tac x=1 in allE)
-      apply auto
-      done
-  next
-    assume as: "x$2 = 1"
-    then have *: "g (x $ 2) $ 2 = 1"
-      using assms(8) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(1)[unfolded subset_eq,rule_format])
-      apply auto
-      done
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
-      apply (erule_tac x=2 in allE)
-      apply auto
-      done
-  next
-    assume as: "x$2 = -1"
-    then have *: "g (x $ 2) $ 2 = - 1"
-      using assms(7) by auto
-    have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0"
-      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
-      unfolding as negatex_def vector_2
-      by auto
-    moreover
-    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
-      apply -
-      apply (rule assms(1)[unfolded subset_eq,rule_format])
-      apply auto
-      done
-    ultimately show False
-      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
-      apply (erule_tac x=2 in allE)
-      apply auto
-      done
-  qed auto
-qed
-
-lemma fashoda_unit_path:
-  fixes f g :: "real \<Rightarrow> real^2"
-  assumes "path f"
-    and "path g"
-    and "path_image f \<subseteq> cbox (-1) 1"
-    and "path_image g \<subseteq> cbox (-1) 1"
-    and "(pathstart f)$1 = -1"
-    and "(pathfinish f)$1 = 1"
-    and "(pathstart g)$2 = -1"
-    and "(pathfinish g)$2 = 1"
-  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof -
-  note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
-  define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real
-  have isc: "iscale ` {- 1..1} \<subseteq> {0..1}"
-    unfolding iscale_def by auto
-  have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t"
-  proof (rule fashoda_unit)
-    show "(f \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1" "(g \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1"
-      using isc and assms(3-4) by (auto simp add: image_comp [symmetric])
-    have *: "continuous_on {- 1..1} iscale"
-      unfolding iscale_def by (rule continuous_intros)+
-    show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)"
-      apply -
-      apply (rule_tac[!] continuous_on_compose[OF *])
-      apply (rule_tac[!] continuous_on_subset[OF _ isc])
-      apply (rule assms)+
-      done
-    have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1"
-      unfolding vec_eq_iff by auto
-    show "(f \<circ> iscale) (- 1) $ 1 = - 1"
-      and "(f \<circ> iscale) 1 $ 1 = 1"
-      and "(g \<circ> iscale) (- 1) $ 2 = -1"
-      and "(g \<circ> iscale) 1 $ 2 = 1"
-      unfolding o_def iscale_def
-      using assms
-      by (auto simp add: *)
-  qed
-  then obtain s t where st:
-      "s \<in> {- 1..1}"
-      "t \<in> {- 1..1}"
-      "(f \<circ> iscale) s = (g \<circ> iscale) t"
-    by auto
-  show thesis
-    apply (rule_tac z = "f (iscale s)" in that)
-    using st
-    unfolding o_def path_image_def image_iff
-    apply -
-    apply (rule_tac x="iscale s" in bexI)
-    prefer 3
-    apply (rule_tac x="iscale t" in bexI)
-    using isc[unfolded subset_eq, rule_format]
-    apply auto
-    done
-qed
-
-lemma fashoda:
-  fixes b :: "real^2"
-  assumes "path f"
-    and "path g"
-    and "path_image f \<subseteq> cbox a b"
-    and "path_image g \<subseteq> cbox a b"
-    and "(pathstart f)$1 = a$1"
-    and "(pathfinish f)$1 = b$1"
-    and "(pathstart g)$2 = a$2"
-    and "(pathfinish g)$2 = b$2"
-  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof -
-  fix P Q S
-  presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis"
-  then show thesis
-    by auto
-next
-  have "cbox a b \<noteq> {}"
-    using assms(3) using path_image_nonempty[of f] by auto
-  then have "a \<le> b"
-    unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less)
-  then show "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)"
-    unfolding less_eq_vec_def forall_2 by auto
-next
-  assume as: "a$1 = b$1"
-  have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2"
-    apply (rule connected_ivt_component_cart)
-    apply (rule connected_path_image assms)+
-    apply (rule pathstart_in_path_image)
-    apply (rule pathfinish_in_path_image)
-    unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"]
-    unfolding pathstart_def
-    apply (auto simp add: less_eq_vec_def mem_interval_cart)
-    done
-  then obtain z :: "real^2" where z: "z \<in> path_image g" "z $ 2 = pathstart f $ 2" ..
-  have "z \<in> cbox a b"
-    using z(1) assms(4)
-    unfolding path_image_def
-    by blast
-  then have "z = f 0"
-    unfolding vec_eq_iff forall_2
-    unfolding z(2) pathstart_def
-    using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1]
-    unfolding mem_interval_cart
-    apply (erule_tac x=1 in allE)
-    using as
-    apply auto
-    done
-  then show thesis
-    apply -
-    apply (rule that[OF _ z(1)])
-    unfolding path_image_def
-    apply auto
-    done
-next
-  assume as: "a$2 = b$2"
-  have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1"
-    apply (rule connected_ivt_component_cart)
-    apply (rule connected_path_image assms)+
-    apply (rule pathstart_in_path_image)
-    apply (rule pathfinish_in_path_image)
-    unfolding assms
-    using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"]
-    unfolding pathstart_def
-    apply (auto simp add: less_eq_vec_def mem_interval_cart)
-    done
-  then obtain z where z: "z \<in> path_image f" "z $ 1 = pathstart g $ 1" ..
-  have "z \<in> cbox a b"
-    using z(1) assms(3)
-    unfolding path_image_def
-    by blast
-  then have "z = g 0"
-    unfolding vec_eq_iff forall_2
-    unfolding z(2) pathstart_def
-    using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2]
-    unfolding mem_interval_cart
-    apply (erule_tac x=2 in allE)
-    using as
-    apply auto
-    done
-  then show thesis
-    apply -
-    apply (rule that[OF z(1)])
-    unfolding path_image_def
-    apply auto
-    done
-next
-  assume as: "a $ 1 < b $ 1 \<and> a $ 2 < b $ 2"
-  have int_nem: "cbox (-1) (1::real^2) \<noteq> {}"
-    unfolding interval_eq_empty_cart by auto
-  obtain z :: "real^2" where z:
-      "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
-      "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
-    apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"]) 
-    unfolding path_def path_image_def pathstart_def pathfinish_def
-    apply (rule_tac[1-2] continuous_on_compose)
-    apply (rule assms[unfolded path_def] continuous_on_interval_bij)+
-    unfolding subset_eq
-    apply(rule_tac[1-2] ballI)
-  proof -
-    fix x
-    assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
-    then obtain y where y:
-        "y \<in> {0..1}"
-        "x = (interval_bij (a, b) (- 1, 1) \<circ> f) y"
-      unfolding image_iff ..
-    show "x \<in> cbox (- 1) 1"
-      unfolding y o_def
-      apply (rule in_interval_interval_bij)
-      using y(1)
-      using assms(3)[unfolded path_image_def subset_eq] int_nem
-      apply auto
-      done
-  next
-    fix x
-    assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
-    then obtain y where y:
-        "y \<in> {0..1}"
-        "x = (interval_bij (a, b) (- 1, 1) \<circ> g) y"
-      unfolding image_iff ..
-    show "x \<in> cbox (- 1) 1"
-      unfolding y o_def
-      apply (rule in_interval_interval_bij)
-      using y(1)
-      using assms(4)[unfolded path_image_def subset_eq] int_nem
-      apply auto
-      done
-  next
-    show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1"
-      and "(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1"
-      and "(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1"
-      and "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1"
-      using assms as
-      by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def)
-         (simp_all add: inner_axis)
-  qed
-  from z(1) obtain zf where zf:
-      "zf \<in> {0..1}"
-      "z = (interval_bij (a, b) (- 1, 1) \<circ> f) zf"
-    unfolding image_iff ..
-  from z(2) obtain zg where zg:
-      "zg \<in> {0..1}"
-      "z = (interval_bij (a, b) (- 1, 1) \<circ> g) zg"
-    unfolding image_iff ..
-  have *: "\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i"
-    unfolding forall_2
-    using as
-    by auto
-  show thesis
-    apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
-    apply (subst zf)
-    defer
-    apply (subst zg)
-    unfolding o_def interval_bij_bij_cart[OF *] path_image_def
-    using zf(1) zg(1)
-    apply auto
-    done
-qed
-
-
-subsection \<open>Some slightly ad hoc lemmas I use below\<close>
-
-lemma segment_vertical:
-  fixes a :: "real^2"
-  assumes "a$1 = b$1"
-  shows "x \<in> closed_segment a b \<longleftrightarrow>
-    x$1 = a$1 \<and> x$1 = b$1 \<and> (a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2)"
-  (is "_ = ?R")
-proof -
-  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
-  {
-    presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L"
-    then show ?thesis
-      unfolding closed_segment_def mem_Collect_eq
-      unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps
-      by blast
-  }
-  {
-    assume ?L
-    then obtain u where u:
-        "x $ 1 = (1 - u) * a $ 1 + u * b $ 1"
-        "x $ 2 = (1 - u) * a $ 2 + u * b $ 2"
-        "0 \<le> u"
-        "u \<le> 1"
-      by blast
-    { fix b a
-      assume "b + u * a > a + u * b"
-      then have "(1 - u) * b > (1 - u) * a"
-        by (auto simp add:field_simps)
-      then have "b \<ge> a"
-        apply (drule_tac mult_left_less_imp_less)
-        using u
-        apply auto
-        done
-      then have "u * a \<le> u * b"
-        apply -
-        apply (rule mult_left_mono[OF _ u(3)]) 
-        using u(3-4)
-        apply (auto simp add: field_simps)
-        done
-    } note * = this
-    {
-      fix a b
-      assume "u * b > u * a"
-      then have "(1 - u) * a \<le> (1 - u) * b"
-        apply -
-        apply (rule mult_left_mono)
-        apply (drule mult_left_less_imp_less)
-        using u
-        apply auto
-        done
-      then have "a + u * b \<le> b + u * a"
-        by (auto simp add: field_simps)
-    } note ** = this
-    then show ?R
-      unfolding u assms
-      using u
-      by (auto simp add:field_simps not_le intro: * **)
-  }
-  {
-    assume ?R
-    then show ?L
-    proof (cases "x$2 = b$2")
-      case True
-      then show ?L
-        apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI)
-        unfolding assms True
-        using \<open>?R\<close>
-        apply (auto simp add: field_simps)
-        done
-    next
-      case False
-      then show ?L
-        apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI)
-        unfolding assms
-        using \<open>?R\<close>
-        apply (auto simp add: field_simps)
-        done
-    qed
-  }
-qed
-
-lemma segment_horizontal:
-  fixes a :: "real^2"
-  assumes "a$2 = b$2"
-  shows "x \<in> closed_segment a b \<longleftrightarrow>
-    x$2 = a$2 \<and> x$2 = b$2 \<and> (a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1)"
-  (is "_ = ?R")
-proof -
-  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
-  {
-    presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L"
-    then show ?thesis
-      unfolding closed_segment_def mem_Collect_eq
-      unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps
-      by blast
-  }
-  {
-    assume ?L
-    then obtain u where u:
-        "x $ 1 = (1 - u) * a $ 1 + u * b $ 1"
-        "x $ 2 = (1 - u) * a $ 2 + u * b $ 2"
-        "0 \<le> u"
-        "u \<le> 1"
-      by blast
-    {
-      fix b a
-      assume "b + u * a > a + u * b"
-      then have "(1 - u) * b > (1 - u) * a"
-        by (auto simp add: field_simps)
-      then have "b \<ge> a"
-        apply (drule_tac mult_left_less_imp_less)
-        using u
-        apply auto
-        done
-      then have "u * a \<le> u * b"
-        apply -
-        apply (rule mult_left_mono[OF _ u(3)])
-        using u(3-4)
-        apply (auto simp add: field_simps)
-        done
-    } note * = this
-    {
-      fix a b
-      assume "u * b > u * a"
-      then have "(1 - u) * a \<le> (1 - u) * b"
-        apply -
-        apply (rule mult_left_mono)
-        apply (drule mult_left_less_imp_less)
-        using u
-        apply auto
-        done
-      then have "a + u * b \<le> b + u * a"
-        by (auto simp add: field_simps)
-    } note ** = this
-    then show ?R
-      unfolding u assms
-      using u
-      by (auto simp add: field_simps not_le intro: * **)
-  }
-  {
-    assume ?R
-    then show ?L
-    proof (cases "x$1 = b$1")
-      case True
-      then show ?L
-        apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI)
-        unfolding assms True
-        using \<open>?R\<close>
-        apply (auto simp add: field_simps)
-        done
-    next
-      case False
-      then show ?L
-        apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI)
-        unfolding assms
-        using \<open>?R\<close>
-        apply (auto simp add: field_simps)
-        done
-    qed
-  }
-qed
-
-
-subsection \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close>
-
-lemma fashoda_interlace:
-  fixes a :: "real^2"
-  assumes "path f"
-    and "path g"
-    and "path_image f \<subseteq> cbox a b"
-    and "path_image g \<subseteq> cbox a b"
-    and "(pathstart f)$2 = a$2"
-    and "(pathfinish f)$2 = a$2"
-    and "(pathstart g)$2 = a$2"
-    and "(pathfinish g)$2 = a$2"
-    and "(pathstart f)$1 < (pathstart g)$1"
-    and "(pathstart g)$1 < (pathfinish f)$1"
-    and "(pathfinish f)$1 < (pathfinish g)$1"
-  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
-proof -
-  have "cbox a b \<noteq> {}"
-    using path_image_nonempty[of f] using assms(3) by auto
-  note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less]
-  have "pathstart f \<in> cbox a b"
-    and "pathfinish f \<in> cbox a b"
-    and "pathstart g \<in> cbox a b"
-    and "pathfinish g \<in> cbox a b"
-    using pathstart_in_path_image pathfinish_in_path_image
-    using assms(3-4)
-    by auto
-  note startfin = this[unfolded mem_interval_cart forall_2]
-  let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++
-     linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++
-     linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++
-     linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])" 
-  let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++
-     linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++
-     linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++
-     linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])"
-  let ?a = "vector[a$1 - 2, a$2 - 3]"
-  let ?b = "vector[b$1 + 2, b$2 + 3]"
-  have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union>
-      path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union>
-      path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union>
-      path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))"
-    "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union>
-      path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>
-      path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>
-      path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)
-      by(auto simp add: path_image_join path_linepath)
-  have abab: "cbox a b \<subseteq> cbox ?a ?b"
-    unfolding interval_cbox_cart[symmetric]
-    by (auto simp add:less_eq_vec_def forall_2 vector_2)
-  obtain z where
-    "z \<in> path_image
-          (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++
-           linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++
-           f +++
-           linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++
-           linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))"
-    "z \<in> path_image
-          (linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++
-           g +++
-           linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++
-           linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++
-           linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))"
-    apply (rule fashoda[of ?P1 ?P2 ?a ?b])
-    unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2
-  proof -
-    show "path ?P1" and "path ?P2"
-      using assms by auto
-    have "path_image ?P1 \<subseteq> cbox ?a ?b"
-      unfolding P1P2 path_image_linepath
-      apply (rule Un_least)+
-      defer 3
-      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_interval_cart forall_2 vector_2
-      using ab startfin abab assms(3)
-      using assms(9-)
-      unfolding assms
-      apply (auto simp add: field_simps box_def)
-      done
-    then show "path_image ?P1 \<subseteq> cbox ?a ?b" .
-    have "path_image ?P2 \<subseteq> cbox ?a ?b"
-      unfolding P1P2 path_image_linepath
-      apply (rule Un_least)+
-      defer 2
-      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
-      unfolding mem_interval_cart forall_2 vector_2
-      using ab startfin abab assms(4)
-      using assms(9-)
-      unfolding assms
-      apply (auto simp add: field_simps box_def)
-      done
-    then show "path_image ?P2 \<subseteq> cbox ?a ?b" .
-    show "a $ 1 - 2 = a $ 1 - 2"
-      and "b $ 1 + 2 = b $ 1 + 2"
-      and "pathstart g $ 2 - 3 = a $ 2 - 3"
-      and "b $ 2 + 3 = b $ 2 + 3"
-      by (auto simp add: assms)
-  qed
-  note z=this[unfolded P1P2 path_image_linepath]
-  show thesis
-    apply (rule that[of z])
-  proof -
-    have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or>
-      z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or>
-      z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or>
-      z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow>
-    (((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or>
-      z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or>
-      z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or>
-      z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False"
-    proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases)
-      case prems: 1
-      have "pathfinish f \<in> cbox a b"
-        using assms(3) pathfinish_in_path_image[of f] by auto 
-      then have "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False"
-        unfolding mem_interval_cart forall_2 by auto
-      then have "z$1 \<noteq> pathfinish f$1"
-        using prems(2)
-        using assms ab
-        by (auto simp add: field_simps)
-      moreover have "pathstart f \<in> cbox a b"
-        using assms(3) pathstart_in_path_image[of f]
-        by auto
-      then have "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False"
-        unfolding mem_interval_cart forall_2
-        by auto
-      then have "z$1 \<noteq> pathstart f$1"
-        using prems(2) using assms ab
-        by (auto simp add: field_simps)
-      ultimately have *: "z$2 = a$2 - 2"
-        using prems(1)
-        by auto
-      have "z$1 \<noteq> pathfinish g$1"
-        using prems(2)
-        using assms ab
-        by (auto simp add: field_simps *)
-      moreover have "pathstart g \<in> cbox a b"
-        using assms(4) pathstart_in_path_image[of g]
-        by auto 
-      note this[unfolded mem_interval_cart forall_2]
-      then have "z$1 \<noteq> pathstart g$1"
-        using prems(1)
-        using assms ab
-        by (auto simp add: field_simps *)
-      ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1"
-        using prems(2)
-        unfolding * assms
-        by (auto simp add: field_simps)
-      then show False
-        unfolding * using ab by auto
-    qed
-    then have "z \<in> path_image f \<or> z \<in> path_image g"
-      using z unfolding Un_iff by blast
-    then have z': "z \<in> cbox a b"
-      using assms(3-4)
-      by auto
-    have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow>
-      z = pathstart f \<or> z = pathfinish f"
-      unfolding vec_eq_iff forall_2 assms
-      by auto
-    with z' show "z \<in> path_image f"
-      using z(1)
-      unfolding Un_iff mem_interval_cart forall_2
-      apply -
-      apply (simp only: segment_vertical segment_horizontal vector_2)
-      unfolding assms
-      apply auto
-      done
-    have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow>
-      z = pathstart g \<or> z = pathfinish g"
-      unfolding vec_eq_iff forall_2 assms
-      by auto
-    with z' show "z \<in> path_image g"
-      using z(2)
-      unfolding Un_iff mem_interval_cart forall_2
-      apply (simp only: segment_vertical segment_horizontal vector_2)
-      unfolding assms
-      apply auto
-      done
-  qed
-qed
-
-(** The Following still needs to be translated. Maybe I will do that later.
-
-(* ------------------------------------------------------------------------- *)
-(* Complement in dimension N >= 2 of set homeomorphic to any interval in     *)
-(* any dimension is (path-)connected. This naively generalizes the argument  *)
-(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer        *)
-(* fixed point theorem", American Mathematical Monthly 1984.                 *)
-(* ------------------------------------------------------------------------- *)
-
-let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove
- (`!p:real^M->real^N a b.
-        ~(interval[a,b] = {}) /\
-        p continuous_on interval[a,b] /\
-        (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y)
-        ==> ?f. f continuous_on (:real^N) /\
-                IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\
-                (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`,
-  REPEAT STRIP_TAC THEN
-  FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN
-  DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN
-  SUBGOAL_THEN `(q:real^N->real^M) continuous_on
-                (IMAGE p (interval[a:real^M,b]))`
-  ASSUME_TAC THENL
-   [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL];
-    ALL_TAC] THEN
-  MP_TAC(ISPECL [`q:real^N->real^M`;
-                 `IMAGE (p:real^M->real^N)
-                 (interval[a,b])`;
-                 `a:real^M`; `b:real^M`]
-        TIETZE_CLOSED_INTERVAL) THEN
-  ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE;
-               COMPACT_IMP_CLOSED] THEN
-  ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
-  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN
-  EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN
-  REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN
-  CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN
-  MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ]
-        CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);;
-
-let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
- (`!s:real^N->bool a b:real^M.
-        s homeomorphic (interval[a,b])
-        ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`,
-  REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN
-  REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
-  MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN
-  DISCH_TAC THEN
-  SUBGOAL_THEN
-   `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\
-          (p:real^M->real^N) x = p y ==> x = y`
-  ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN
-  FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN
-  DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN
-  ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
-  ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV;
-                  NOT_BOUNDED_UNIV] THEN
-  ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN
-  X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN
-  SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
-  SUBGOAL_THEN `bounded((path_component s c) UNION
-                        (IMAGE (p:real^M->real^N) (interval[a,b])))`
-  MP_TAC THENL
-   [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED;
-                 COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-    ALL_TAC] THEN
-  DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN
-  REWRITE_TAC[UNION_SUBSET] THEN
-  DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
-  MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`]
-    RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN
-  ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN
-  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN
-  DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC
-   (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
-  REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN
-  ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN
-  SUBGOAL_THEN
-    `(q:real^N->real^N) continuous_on
-     (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))`
-  MP_TAC THENL
-   [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN
-    REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN
-    REPEAT CONJ_TAC THENL
-     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
-      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
-                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-      ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV];
-      ALL_TAC] THEN
-    X_GEN_TAC `z:real^N` THEN
-    REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN
-    STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
-    MP_TAC(ISPECL
-     [`path_component s (z:real^N)`; `path_component s (c:real^N)`]
-     OPEN_INTER_CLOSURE_EQ_EMPTY) THEN
-    ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL
-     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
-      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
-                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
-      REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN
-      DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN
-      GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN
-      REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]];
-    ALL_TAC] THEN
-  SUBGOAL_THEN
-   `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) =
-    (:real^N)`
-  SUBST1_TAC THENL
-   [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN
-    REWRITE_TAC[CLOSURE_SUBSET];
-    DISCH_TAC] THEN
-  MP_TAC(ISPECL
-   [`(\x. &2 % c - x) o
-     (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`;
-    `cball(c:real^N,B)`]
-    BROUWER) THEN
-  REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN
-  ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN
-  SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL
-   [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN
-    REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN
-    ASM SET_TAC[PATH_COMPONENT_REFL_EQ];
-    ALL_TAC] THEN
-  REPEAT CONJ_TAC THENL
-   [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
-    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL
-     [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_MUL THEN
-    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
-    REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN
-    MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN
-    ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-    SUBGOAL_THEN
-     `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)`
-    SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN
-    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
-    ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST;
-                 CONTINUOUS_ON_LIFT_NORM];
-    REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN
-    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
-    REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN
-    REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
-    ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-    ASM_REAL_ARITH_TAC;
-    REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN
-    REWRITE_TAC[IN_CBALL; o_THM; dist] THEN
-    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
-    REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN
-    ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL
-     [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN
-      REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
-      ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
-      ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN
-      UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN
-      REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB];
-      EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN
-      REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN
-      ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN
-      SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL
-       [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN
-      ASM_REWRITE_TAC[] THEN
-      MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN
-      ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);;
-
-let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
- (`!s:real^N->bool a b:real^M.
-        2 <= dimindex(:N) /\ s homeomorphic interval[a,b]
-        ==> path_connected((:real^N) DIFF s)`,
-  REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
-  FIRST_ASSUM(MP_TAC o MATCH_MP
-    UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
-  ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN
-  ABBREV_TAC `t = (:real^N) DIFF s` THEN
-  DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN
-  STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN
-  REWRITE_TAC[COMPACT_INTERVAL] THEN
-  DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN
-  REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN
-  X_GEN_TAC `B:real` THEN STRIP_TAC THEN
-  SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\
-                (?v:real^N. v IN path_component t y /\ B < norm(v))`
-  STRIP_ASSUME_TAC THENL
-   [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN
-  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_SYM THEN
-  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN
-  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
-  MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN
-  EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL
-   [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE
-     `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN
-    ASM_REWRITE_TAC[SUBSET; IN_CBALL_0];
-    MP_TAC(ISPEC `cball(vec 0:real^N,B)`
-       PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN
-    ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN
-    REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
-    DISCH_THEN MATCH_MP_TAC THEN
-    ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);;
-
-(* ------------------------------------------------------------------------- *)
-(* In particular, apply all these to the special case of an arc.             *)
-(* ------------------------------------------------------------------------- *)
-
-let RETRACTION_ARC = prove
- (`!p. arc p
-       ==> ?f. f continuous_on (:real^N) /\
-               IMAGE f (:real^N) SUBSET path_image p /\
-               (!x. x IN path_image p ==> f x = x)`,
-  REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN
-  MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN
-  ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;
-
-let PATH_CONNECTED_ARC_COMPLEMENT = prove
- (`!p. 2 <= dimindex(:N) /\ arc p
-       ==> path_connected((:real^N) DIFF path_image p)`,
-  REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN
-  MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`]
-    PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
-  ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN
-  ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN
-  MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN
-  EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);;
-
-let CONNECTED_ARC_COMPLEMENT = prove
- (`!p. 2 <= dimindex(:N) /\ arc p
-       ==> connected((:real^N) DIFF path_image p)`,
-  SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *)
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Fashoda_Theorem.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,1136 @@
+(*  Author:     John Harrison
+    Author:     Robert Himmelmann, TU Muenchen (translation from HOL light)
+*)
+
+section \<open>Fashoda meet theorem\<close>
+
+theory Fashoda_Theorem
+imports Brouwer_Fixpoint Path_Connected Cartesian_Euclidean_Space
+begin
+
+subsection \<open>Bijections between intervals.\<close>
+
+definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
+  where "interval_bij =
+    (\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))"
+
+lemma interval_bij_affine:
+  "interval_bij (a,b) (u,v) = (\<lambda>x. (\<Sum>i\<in>Basis. ((v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (x\<bullet>i)) *\<^sub>R i) +
+    (\<Sum>i\<in>Basis. (u\<bullet>i - (v\<bullet>i - u\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (a\<bullet>i)) *\<^sub>R i))"
+  by (auto simp: setsum.distrib[symmetric] scaleR_add_left[symmetric] interval_bij_def fun_eq_iff
+    field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum.cong)
+
+lemma continuous_interval_bij:
+  fixes a b :: "'a::euclidean_space"
+  shows "continuous (at x) (interval_bij (a, b) (u, v))"
+  by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros)
+
+lemma continuous_on_interval_bij: "continuous_on s (interval_bij (a, b) (u, v))"
+  apply(rule continuous_at_imp_continuous_on)
+  apply (rule, rule continuous_interval_bij)
+  done
+
+lemma in_interval_interval_bij:
+  fixes a b u v x :: "'a::euclidean_space"
+  assumes "x \<in> cbox a b"
+    and "cbox u v \<noteq> {}"
+  shows "interval_bij (a, b) (u, v) x \<in> cbox u v"
+  apply (simp only: interval_bij_def split_conv mem_box inner_setsum_left_Basis cong: ball_cong)
+  apply safe
+proof -
+  fix i :: 'a
+  assume i: "i \<in> Basis"
+  have "cbox a b \<noteq> {}"
+    using assms by auto
+  with i have *: "a\<bullet>i \<le> b\<bullet>i" "u\<bullet>i \<le> v\<bullet>i"
+    using assms(2) by (auto simp add: box_eq_empty)
+  have x: "a\<bullet>i\<le>x\<bullet>i" "x\<bullet>i\<le>b\<bullet>i"
+    using assms(1)[unfolded mem_box] using i by auto
+  have "0 \<le> (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
+    using * x by auto
+  then show "u \<bullet> i \<le> u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i)"
+    using * by auto
+  have "((x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i)) * (v \<bullet> i - u \<bullet> i) \<le> 1 * (v \<bullet> i - u \<bullet> i)"
+    apply (rule mult_right_mono)
+    unfolding divide_le_eq_1
+    using * x
+    apply auto
+    done
+  then show "u \<bullet> i + (x \<bullet> i - a \<bullet> i) / (b \<bullet> i - a \<bullet> i) * (v \<bullet> i - u \<bullet> i) \<le> v \<bullet> i"
+    using * by auto
+qed
+
+lemma interval_bij_bij:
+  "\<forall>(i::'a::euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow>
+    interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x"
+  by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a])
+
+lemma interval_bij_bij_cart: fixes x::"real^'n" assumes "\<forall>i. a$i < b$i \<and> u$i < v$i"
+  shows "interval_bij (a,b) (u,v) (interval_bij (u,v) (a,b) x) = x"
+  using assms by (intro interval_bij_bij) (auto simp: Basis_vec_def inner_axis)
+
+
+subsection \<open>Fashoda meet theorem\<close>
+
+lemma infnorm_2:
+  fixes x :: "real^2"
+  shows "infnorm x = max \<bar>x$1\<bar> \<bar>x$2\<bar>"
+  unfolding infnorm_cart UNIV_2 by (rule cSup_eq) auto
+
+lemma infnorm_eq_1_2:
+  fixes x :: "real^2"
+  shows "infnorm x = 1 \<longleftrightarrow>
+    \<bar>x$1\<bar> \<le> 1 \<and> \<bar>x$2\<bar> \<le> 1 \<and> (x$1 = -1 \<or> x$1 = 1 \<or> x$2 = -1 \<or> x$2 = 1)"
+  unfolding infnorm_2 by auto
+
+lemma infnorm_eq_1_imp:
+  fixes x :: "real^2"
+  assumes "infnorm x = 1"
+  shows "\<bar>x$1\<bar> \<le> 1" and "\<bar>x$2\<bar> \<le> 1"
+  using assms unfolding infnorm_eq_1_2 by auto
+
+lemma fashoda_unit:
+  fixes f g :: "real \<Rightarrow> real^2"
+  assumes "f ` {-1 .. 1} \<subseteq> cbox (-1) 1"
+    and "g ` {-1 .. 1} \<subseteq> cbox (-1) 1"
+    and "continuous_on {-1 .. 1} f"
+    and "continuous_on {-1 .. 1} g"
+    and "f (- 1)$1 = - 1"
+    and "f 1$1 = 1" "g (- 1) $2 = -1"
+    and "g 1 $2 = 1"
+  shows "\<exists>s\<in>{-1 .. 1}. \<exists>t\<in>{-1 .. 1}. f s = g t"
+proof (rule ccontr)
+  assume "\<not> ?thesis"
+  note as = this[unfolded bex_simps,rule_format]
+  define sqprojection
+    where [abs_def]: "sqprojection z = (inverse (infnorm z)) *\<^sub>R z" for z :: "real^2"
+  define negatex :: "real^2 \<Rightarrow> real^2"
+    where "negatex x = (vector [-(x$1), x$2])" for x
+  have lem1: "\<forall>z::real^2. infnorm (negatex z) = infnorm z"
+    unfolding negatex_def infnorm_2 vector_2 by auto
+  have lem2: "\<forall>z. z \<noteq> 0 \<longrightarrow> infnorm (sqprojection z) = 1"
+    unfolding sqprojection_def
+    unfolding infnorm_mul[unfolded scalar_mult_eq_scaleR]
+    unfolding abs_inverse real_abs_infnorm
+    apply (subst infnorm_eq_0[symmetric])
+    apply auto
+    done
+  let ?F = "\<lambda>w::real^2. (f \<circ> (\<lambda>x. x$1)) w - (g \<circ> (\<lambda>x. x$2)) w"
+  have *: "\<And>i. (\<lambda>x::real^2. x $ i) ` cbox (- 1) 1 = {-1 .. 1}"
+    apply (rule set_eqI)
+    unfolding image_iff Bex_def mem_interval_cart interval_cbox_cart
+    apply rule
+    defer
+    apply (rule_tac x="vec x" in exI)
+    apply auto
+    done
+  {
+    fix x
+    assume "x \<in> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w) ` (cbox (- 1) (1::real^2))"
+    then obtain w :: "real^2" where w:
+        "w \<in> cbox (- 1) 1"
+        "x = (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w"
+      unfolding image_iff ..
+    then have "x \<noteq> 0"
+      using as[of "w$1" "w$2"]
+      unfolding mem_interval_cart atLeastAtMost_iff
+      by auto
+  } note x0 = this
+  have 21: "\<And>i::2. i \<noteq> 1 \<Longrightarrow> i = 2"
+    using UNIV_2 by auto
+  have 1: "box (- 1) (1::real^2) \<noteq> {}"
+    unfolding interval_eq_empty_cart by auto
+  have 2: "continuous_on (cbox (- 1) 1) (negatex \<circ> sqprojection \<circ> ?F)"
+    apply (intro continuous_intros continuous_on_component)
+    unfolding *
+    apply (rule assms)+
+    apply (subst sqprojection_def)
+    apply (intro continuous_intros)
+    apply (simp add: infnorm_eq_0 x0)
+    apply (rule linear_continuous_on)
+  proof -
+    show "bounded_linear negatex"
+      apply (rule bounded_linearI')
+      unfolding vec_eq_iff
+    proof (rule_tac[!] allI)
+      fix i :: 2
+      fix x y :: "real^2"
+      fix c :: real
+      show "negatex (x + y) $ i =
+        (negatex x + negatex y) $ i" "negatex (c *\<^sub>R x) $ i = (c *\<^sub>R negatex x) $ i"
+        apply -
+        apply (case_tac[!] "i\<noteq>1")
+        prefer 3
+        apply (drule_tac[1-2] 21)
+        unfolding negatex_def
+        apply (auto simp add:vector_2)
+        done
+    qed
+  qed
+  have 3: "(negatex \<circ> sqprojection \<circ> ?F) ` cbox (-1) 1 \<subseteq> cbox (-1) 1"
+    unfolding subset_eq
+  proof (rule, goal_cases)
+    case (1 x)
+    then obtain y :: "real^2" where y:
+        "y \<in> cbox (- 1) 1"
+        "x = (negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) y"
+      unfolding image_iff ..
+    have "?F y \<noteq> 0"
+      apply (rule x0)
+      using y(1)
+      apply auto
+      done
+    then have *: "infnorm (sqprojection (?F y)) = 1"
+      unfolding y o_def
+      by - (rule lem2[rule_format])
+    have "infnorm x = 1"
+      unfolding *[symmetric] y o_def
+      by (rule lem1[rule_format])
+    then show "x \<in> cbox (-1) 1"
+      unfolding mem_interval_cart interval_cbox_cart infnorm_2
+      apply -
+      apply rule
+    proof -
+      fix i
+      assume "max \<bar>x $ 1\<bar> \<bar>x $ 2\<bar> = 1"
+      then show "(- 1) $ i \<le> x $ i \<and> x $ i \<le> 1 $ i"
+        apply (cases "i = 1")
+        defer
+        apply (drule 21)
+        apply auto
+        done
+    qed
+  qed
+  obtain x :: "real^2" where x:
+      "x \<in> cbox (- 1) 1"
+      "(negatex \<circ> sqprojection \<circ> (\<lambda>w. (f \<circ> (\<lambda>x. x $ 1)) w - (g \<circ> (\<lambda>x. x $ 2)) w)) x = x"
+    apply (rule brouwer_weak[of "cbox (- 1) (1::real^2)" "negatex \<circ> sqprojection \<circ> ?F"])
+    apply (rule compact_cbox convex_box)+
+    unfolding interior_cbox
+    apply (rule 1 2 3)+
+    apply blast
+    done
+  have "?F x \<noteq> 0"
+    apply (rule x0)
+    using x(1)
+    apply auto
+    done
+  then have *: "infnorm (sqprojection (?F x)) = 1"
+    unfolding o_def
+    by (rule lem2[rule_format])
+  have nx: "infnorm x = 1"
+    apply (subst x(2)[symmetric])
+    unfolding *[symmetric] o_def
+    apply (rule lem1[rule_format])
+    done
+  have "\<forall>x i. x \<noteq> 0 \<longrightarrow> (0 < (sqprojection x)$i \<longleftrightarrow> 0 < x$i)"
+    and "\<forall>x i. x \<noteq> 0 \<longrightarrow> ((sqprojection x)$i < 0 \<longleftrightarrow> x$i < 0)"
+    apply -
+    apply (rule_tac[!] allI impI)+
+  proof -
+    fix x :: "real^2"
+    fix i :: 2
+    assume x: "x \<noteq> 0"
+    have "inverse (infnorm x) > 0"
+      using x[unfolded infnorm_pos_lt[symmetric]] by auto
+    then show "(0 < sqprojection x $ i) = (0 < x $ i)"
+      and "(sqprojection x $ i < 0) = (x $ i < 0)"
+      unfolding sqprojection_def vector_component_simps vector_scaleR_component real_scaleR_def
+      unfolding zero_less_mult_iff mult_less_0_iff
+      by (auto simp add: field_simps)
+  qed
+  note lem3 = this[rule_format]
+  have x1: "x $ 1 \<in> {- 1..1::real}" "x $ 2 \<in> {- 1..1::real}"
+    using x(1) unfolding mem_interval_cart by auto
+  then have nz: "f (x $ 1) - g (x $ 2) \<noteq> 0"
+    unfolding right_minus_eq
+    apply -
+    apply (rule as)
+    apply auto
+    done
+  have "x $ 1 = -1 \<or> x $ 1 = 1 \<or> x $ 2 = -1 \<or> x $ 2 = 1"
+    using nx unfolding infnorm_eq_1_2 by auto
+  then show False
+  proof -
+    fix P Q R S
+    presume "P \<or> Q \<or> R \<or> S"
+      and "P \<Longrightarrow> False"
+      and "Q \<Longrightarrow> False"
+      and "R \<Longrightarrow> False"
+      and "S \<Longrightarrow> False"
+    then show False by auto
+  next
+    assume as: "x$1 = 1"
+    then have *: "f (x $ 1) $ 1 = 1"
+      using assms(6) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 1 < 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
+      unfolding as negatex_def vector_2
+      by auto
+    moreover
+    from x1 have "g (x $ 2) \<in> cbox (-1) 1"
+      apply -
+      apply (rule assms(2)[unfolded subset_eq,rule_format])
+      apply auto
+      done
+    ultimately show False
+      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
+      apply (erule_tac x=1 in allE)
+      apply auto
+      done
+  next
+    assume as: "x$1 = -1"
+    then have *: "f (x $ 1) $ 1 = - 1"
+      using assms(5) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 1 > 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=1]]
+      unfolding as negatex_def vector_2
+      by auto
+    moreover
+    from x1 have "g (x $ 2) \<in> cbox (-1) 1"
+      apply -
+      apply (rule assms(2)[unfolded subset_eq,rule_format])
+      apply auto
+      done
+    ultimately show False
+      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
+      apply (erule_tac x=1 in allE)
+      apply auto
+      done
+  next
+    assume as: "x$2 = 1"
+    then have *: "g (x $ 2) $ 2 = 1"
+      using assms(8) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 2 > 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
+      unfolding as negatex_def vector_2
+      by auto
+    moreover
+    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
+      apply -
+      apply (rule assms(1)[unfolded subset_eq,rule_format])
+      apply auto
+      done
+    ultimately show False
+      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
+      apply (erule_tac x=2 in allE)
+      apply auto
+      done
+  next
+    assume as: "x$2 = -1"
+    then have *: "g (x $ 2) $ 2 = - 1"
+      using assms(7) by auto
+    have "sqprojection (f (x$1) - g (x$2)) $ 2 < 0"
+      using x(2)[unfolded o_def vec_eq_iff,THEN spec[where x=2]]
+      unfolding as negatex_def vector_2
+      by auto
+    moreover
+    from x1 have "f (x $ 1) \<in> cbox (-1) 1"
+      apply -
+      apply (rule assms(1)[unfolded subset_eq,rule_format])
+      apply auto
+      done
+    ultimately show False
+      unfolding lem3[OF nz] vector_component_simps * mem_interval_cart
+      apply (erule_tac x=2 in allE)
+      apply auto
+      done
+  qed auto
+qed
+
+lemma fashoda_unit_path:
+  fixes f g :: "real \<Rightarrow> real^2"
+  assumes "path f"
+    and "path g"
+    and "path_image f \<subseteq> cbox (-1) 1"
+    and "path_image g \<subseteq> cbox (-1) 1"
+    and "(pathstart f)$1 = -1"
+    and "(pathfinish f)$1 = 1"
+    and "(pathstart g)$2 = -1"
+    and "(pathfinish g)$2 = 1"
+  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
+proof -
+  note assms=assms[unfolded path_def pathstart_def pathfinish_def path_image_def]
+  define iscale where [abs_def]: "iscale z = inverse 2 *\<^sub>R (z + 1)" for z :: real
+  have isc: "iscale ` {- 1..1} \<subseteq> {0..1}"
+    unfolding iscale_def by auto
+  have "\<exists>s\<in>{- 1..1}. \<exists>t\<in>{- 1..1}. (f \<circ> iscale) s = (g \<circ> iscale) t"
+  proof (rule fashoda_unit)
+    show "(f \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1" "(g \<circ> iscale) ` {- 1..1} \<subseteq> cbox (- 1) 1"
+      using isc and assms(3-4) by (auto simp add: image_comp [symmetric])
+    have *: "continuous_on {- 1..1} iscale"
+      unfolding iscale_def by (rule continuous_intros)+
+    show "continuous_on {- 1..1} (f \<circ> iscale)" "continuous_on {- 1..1} (g \<circ> iscale)"
+      apply -
+      apply (rule_tac[!] continuous_on_compose[OF *])
+      apply (rule_tac[!] continuous_on_subset[OF _ isc])
+      apply (rule assms)+
+      done
+    have *: "(1 / 2) *\<^sub>R (1 + (1::real^1)) = 1"
+      unfolding vec_eq_iff by auto
+    show "(f \<circ> iscale) (- 1) $ 1 = - 1"
+      and "(f \<circ> iscale) 1 $ 1 = 1"
+      and "(g \<circ> iscale) (- 1) $ 2 = -1"
+      and "(g \<circ> iscale) 1 $ 2 = 1"
+      unfolding o_def iscale_def
+      using assms
+      by (auto simp add: *)
+  qed
+  then obtain s t where st:
+      "s \<in> {- 1..1}"
+      "t \<in> {- 1..1}"
+      "(f \<circ> iscale) s = (g \<circ> iscale) t"
+    by auto
+  show thesis
+    apply (rule_tac z = "f (iscale s)" in that)
+    using st
+    unfolding o_def path_image_def image_iff
+    apply -
+    apply (rule_tac x="iscale s" in bexI)
+    prefer 3
+    apply (rule_tac x="iscale t" in bexI)
+    using isc[unfolded subset_eq, rule_format]
+    apply auto
+    done
+qed
+
+lemma fashoda:
+  fixes b :: "real^2"
+  assumes "path f"
+    and "path g"
+    and "path_image f \<subseteq> cbox a b"
+    and "path_image g \<subseteq> cbox a b"
+    and "(pathstart f)$1 = a$1"
+    and "(pathfinish f)$1 = b$1"
+    and "(pathstart g)$2 = a$2"
+    and "(pathfinish g)$2 = b$2"
+  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
+proof -
+  fix P Q S
+  presume "P \<or> Q \<or> S" "P \<Longrightarrow> thesis" and "Q \<Longrightarrow> thesis" and "S \<Longrightarrow> thesis"
+  then show thesis
+    by auto
+next
+  have "cbox a b \<noteq> {}"
+    using assms(3) using path_image_nonempty[of f] by auto
+  then have "a \<le> b"
+    unfolding interval_eq_empty_cart less_eq_vec_def by (auto simp add: not_less)
+  then show "a$1 = b$1 \<or> a$2 = b$2 \<or> (a$1 < b$1 \<and> a$2 < b$2)"
+    unfolding less_eq_vec_def forall_2 by auto
+next
+  assume as: "a$1 = b$1"
+  have "\<exists>z\<in>path_image g. z$2 = (pathstart f)$2"
+    apply (rule connected_ivt_component_cart)
+    apply (rule connected_path_image assms)+
+    apply (rule pathstart_in_path_image)
+    apply (rule pathfinish_in_path_image)
+    unfolding assms using assms(3)[unfolded path_image_def subset_eq,rule_format,of "f 0"]
+    unfolding pathstart_def
+    apply (auto simp add: less_eq_vec_def mem_interval_cart)
+    done
+  then obtain z :: "real^2" where z: "z \<in> path_image g" "z $ 2 = pathstart f $ 2" ..
+  have "z \<in> cbox a b"
+    using z(1) assms(4)
+    unfolding path_image_def
+    by blast
+  then have "z = f 0"
+    unfolding vec_eq_iff forall_2
+    unfolding z(2) pathstart_def
+    using assms(3)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "f 0" 1]
+    unfolding mem_interval_cart
+    apply (erule_tac x=1 in allE)
+    using as
+    apply auto
+    done
+  then show thesis
+    apply -
+    apply (rule that[OF _ z(1)])
+    unfolding path_image_def
+    apply auto
+    done
+next
+  assume as: "a$2 = b$2"
+  have "\<exists>z\<in>path_image f. z$1 = (pathstart g)$1"
+    apply (rule connected_ivt_component_cart)
+    apply (rule connected_path_image assms)+
+    apply (rule pathstart_in_path_image)
+    apply (rule pathfinish_in_path_image)
+    unfolding assms
+    using assms(4)[unfolded path_image_def subset_eq,rule_format,of "g 0"]
+    unfolding pathstart_def
+    apply (auto simp add: less_eq_vec_def mem_interval_cart)
+    done
+  then obtain z where z: "z \<in> path_image f" "z $ 1 = pathstart g $ 1" ..
+  have "z \<in> cbox a b"
+    using z(1) assms(3)
+    unfolding path_image_def
+    by blast
+  then have "z = g 0"
+    unfolding vec_eq_iff forall_2
+    unfolding z(2) pathstart_def
+    using assms(4)[unfolded path_image_def subset_eq mem_interval_cart,rule_format,of "g 0" 2]
+    unfolding mem_interval_cart
+    apply (erule_tac x=2 in allE)
+    using as
+    apply auto
+    done
+  then show thesis
+    apply -
+    apply (rule that[OF z(1)])
+    unfolding path_image_def
+    apply auto
+    done
+next
+  assume as: "a $ 1 < b $ 1 \<and> a $ 2 < b $ 2"
+  have int_nem: "cbox (-1) (1::real^2) \<noteq> {}"
+    unfolding interval_eq_empty_cart by auto
+  obtain z :: "real^2" where z:
+      "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
+      "z \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
+    apply (rule fashoda_unit_path[of "interval_bij (a,b) (- 1,1) \<circ> f" "interval_bij (a,b) (- 1,1) \<circ> g"])
+    unfolding path_def path_image_def pathstart_def pathfinish_def
+    apply (rule_tac[1-2] continuous_on_compose)
+    apply (rule assms[unfolded path_def] continuous_on_interval_bij)+
+    unfolding subset_eq
+    apply(rule_tac[1-2] ballI)
+  proof -
+    fix x
+    assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> f) ` {0..1}"
+    then obtain y where y:
+        "y \<in> {0..1}"
+        "x = (interval_bij (a, b) (- 1, 1) \<circ> f) y"
+      unfolding image_iff ..
+    show "x \<in> cbox (- 1) 1"
+      unfolding y o_def
+      apply (rule in_interval_interval_bij)
+      using y(1)
+      using assms(3)[unfolded path_image_def subset_eq] int_nem
+      apply auto
+      done
+  next
+    fix x
+    assume "x \<in> (interval_bij (a, b) (- 1, 1) \<circ> g) ` {0..1}"
+    then obtain y where y:
+        "y \<in> {0..1}"
+        "x = (interval_bij (a, b) (- 1, 1) \<circ> g) y"
+      unfolding image_iff ..
+    show "x \<in> cbox (- 1) 1"
+      unfolding y o_def
+      apply (rule in_interval_interval_bij)
+      using y(1)
+      using assms(4)[unfolded path_image_def subset_eq] int_nem
+      apply auto
+      done
+  next
+    show "(interval_bij (a, b) (- 1, 1) \<circ> f) 0 $ 1 = -1"
+      and "(interval_bij (a, b) (- 1, 1) \<circ> f) 1 $ 1 = 1"
+      and "(interval_bij (a, b) (- 1, 1) \<circ> g) 0 $ 2 = -1"
+      and "(interval_bij (a, b) (- 1, 1) \<circ> g) 1 $ 2 = 1"
+      using assms as
+      by (simp_all add: axis_in_Basis cart_eq_inner_axis pathstart_def pathfinish_def interval_bij_def)
+         (simp_all add: inner_axis)
+  qed
+  from z(1) obtain zf where zf:
+      "zf \<in> {0..1}"
+      "z = (interval_bij (a, b) (- 1, 1) \<circ> f) zf"
+    unfolding image_iff ..
+  from z(2) obtain zg where zg:
+      "zg \<in> {0..1}"
+      "z = (interval_bij (a, b) (- 1, 1) \<circ> g) zg"
+    unfolding image_iff ..
+  have *: "\<forall>i. (- 1) $ i < (1::real^2) $ i \<and> a $ i < b $ i"
+    unfolding forall_2
+    using as
+    by auto
+  show thesis
+    apply (rule_tac z="interval_bij (- 1,1) (a,b) z" in that)
+    apply (subst zf)
+    defer
+    apply (subst zg)
+    unfolding o_def interval_bij_bij_cart[OF *] path_image_def
+    using zf(1) zg(1)
+    apply auto
+    done
+qed
+
+
+subsection \<open>Some slightly ad hoc lemmas I use below\<close>
+
+lemma segment_vertical:
+  fixes a :: "real^2"
+  assumes "a$1 = b$1"
+  shows "x \<in> closed_segment a b \<longleftrightarrow>
+    x$1 = a$1 \<and> x$1 = b$1 \<and> (a$2 \<le> x$2 \<and> x$2 \<le> b$2 \<or> b$2 \<le> x$2 \<and> x$2 \<le> a$2)"
+  (is "_ = ?R")
+proof -
+  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
+  {
+    presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L"
+    then show ?thesis
+      unfolding closed_segment_def mem_Collect_eq
+      unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps
+      by blast
+  }
+  {
+    assume ?L
+    then obtain u where u:
+        "x $ 1 = (1 - u) * a $ 1 + u * b $ 1"
+        "x $ 2 = (1 - u) * a $ 2 + u * b $ 2"
+        "0 \<le> u"
+        "u \<le> 1"
+      by blast
+    { fix b a
+      assume "b + u * a > a + u * b"
+      then have "(1 - u) * b > (1 - u) * a"
+        by (auto simp add:field_simps)
+      then have "b \<ge> a"
+        apply (drule_tac mult_left_less_imp_less)
+        using u
+        apply auto
+        done
+      then have "u * a \<le> u * b"
+        apply -
+        apply (rule mult_left_mono[OF _ u(3)])
+        using u(3-4)
+        apply (auto simp add: field_simps)
+        done
+    } note * = this
+    {
+      fix a b
+      assume "u * b > u * a"
+      then have "(1 - u) * a \<le> (1 - u) * b"
+        apply -
+        apply (rule mult_left_mono)
+        apply (drule mult_left_less_imp_less)
+        using u
+        apply auto
+        done
+      then have "a + u * b \<le> b + u * a"
+        by (auto simp add: field_simps)
+    } note ** = this
+    then show ?R
+      unfolding u assms
+      using u
+      by (auto simp add:field_simps not_le intro: * **)
+  }
+  {
+    assume ?R
+    then show ?L
+    proof (cases "x$2 = b$2")
+      case True
+      then show ?L
+        apply (rule_tac x="(x$2 - a$2) / (b$2 - a$2)" in exI)
+        unfolding assms True
+        using \<open>?R\<close>
+        apply (auto simp add: field_simps)
+        done
+    next
+      case False
+      then show ?L
+        apply (rule_tac x="1 - (x$2 - b$2) / (a$2 - b$2)" in exI)
+        unfolding assms
+        using \<open>?R\<close>
+        apply (auto simp add: field_simps)
+        done
+    qed
+  }
+qed
+
+lemma segment_horizontal:
+  fixes a :: "real^2"
+  assumes "a$2 = b$2"
+  shows "x \<in> closed_segment a b \<longleftrightarrow>
+    x$2 = a$2 \<and> x$2 = b$2 \<and> (a$1 \<le> x$1 \<and> x$1 \<le> b$1 \<or> b$1 \<le> x$1 \<and> x$1 \<le> a$1)"
+  (is "_ = ?R")
+proof -
+  let ?L = "\<exists>u. (x $ 1 = (1 - u) * a $ 1 + u * b $ 1 \<and> x $ 2 = (1 - u) * a $ 2 + u * b $ 2) \<and> 0 \<le> u \<and> u \<le> 1"
+  {
+    presume "?L \<Longrightarrow> ?R" and "?R \<Longrightarrow> ?L"
+    then show ?thesis
+      unfolding closed_segment_def mem_Collect_eq
+      unfolding vec_eq_iff forall_2 scalar_mult_eq_scaleR[symmetric] vector_component_simps
+      by blast
+  }
+  {
+    assume ?L
+    then obtain u where u:
+        "x $ 1 = (1 - u) * a $ 1 + u * b $ 1"
+        "x $ 2 = (1 - u) * a $ 2 + u * b $ 2"
+        "0 \<le> u"
+        "u \<le> 1"
+      by blast
+    {
+      fix b a
+      assume "b + u * a > a + u * b"
+      then have "(1 - u) * b > (1 - u) * a"
+        by (auto simp add: field_simps)
+      then have "b \<ge> a"
+        apply (drule_tac mult_left_less_imp_less)
+        using u
+        apply auto
+        done
+      then have "u * a \<le> u * b"
+        apply -
+        apply (rule mult_left_mono[OF _ u(3)])
+        using u(3-4)
+        apply (auto simp add: field_simps)
+        done
+    } note * = this
+    {
+      fix a b
+      assume "u * b > u * a"
+      then have "(1 - u) * a \<le> (1 - u) * b"
+        apply -
+        apply (rule mult_left_mono)
+        apply (drule mult_left_less_imp_less)
+        using u
+        apply auto
+        done
+      then have "a + u * b \<le> b + u * a"
+        by (auto simp add: field_simps)
+    } note ** = this
+    then show ?R
+      unfolding u assms
+      using u
+      by (auto simp add: field_simps not_le intro: * **)
+  }
+  {
+    assume ?R
+    then show ?L
+    proof (cases "x$1 = b$1")
+      case True
+      then show ?L
+        apply (rule_tac x="(x$1 - a$1) / (b$1 - a$1)" in exI)
+        unfolding assms True
+        using \<open>?R\<close>
+        apply (auto simp add: field_simps)
+        done
+    next
+      case False
+      then show ?L
+        apply (rule_tac x="1 - (x$1 - b$1) / (a$1 - b$1)" in exI)
+        unfolding assms
+        using \<open>?R\<close>
+        apply (auto simp add: field_simps)
+        done
+    qed
+  }
+qed
+
+
+subsection \<open>Useful Fashoda corollary pointed out to me by Tom Hales\<close>
+
+lemma fashoda_interlace:
+  fixes a :: "real^2"
+  assumes "path f"
+    and "path g"
+    and "path_image f \<subseteq> cbox a b"
+    and "path_image g \<subseteq> cbox a b"
+    and "(pathstart f)$2 = a$2"
+    and "(pathfinish f)$2 = a$2"
+    and "(pathstart g)$2 = a$2"
+    and "(pathfinish g)$2 = a$2"
+    and "(pathstart f)$1 < (pathstart g)$1"
+    and "(pathstart g)$1 < (pathfinish f)$1"
+    and "(pathfinish f)$1 < (pathfinish g)$1"
+  obtains z where "z \<in> path_image f" and "z \<in> path_image g"
+proof -
+  have "cbox a b \<noteq> {}"
+    using path_image_nonempty[of f] using assms(3) by auto
+  note ab=this[unfolded interval_eq_empty_cart not_ex forall_2 not_less]
+  have "pathstart f \<in> cbox a b"
+    and "pathfinish f \<in> cbox a b"
+    and "pathstart g \<in> cbox a b"
+    and "pathfinish g \<in> cbox a b"
+    using pathstart_in_path_image pathfinish_in_path_image
+    using assms(3-4)
+    by auto
+  note startfin = this[unfolded mem_interval_cart forall_2]
+  let ?P1 = "linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2]) +++
+     linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f) +++ f +++
+     linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2]) +++
+     linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2])"
+  let ?P2 = "linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g) +++ g +++
+     linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1]) +++
+     linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1]) +++
+     linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3])"
+  let ?a = "vector[a$1 - 2, a$2 - 3]"
+  let ?b = "vector[b$1 + 2, b$2 + 3]"
+  have P1P2: "path_image ?P1 = path_image (linepath (vector[a$1 - 2, a$2 - 2]) (vector[(pathstart f)$1,a$2 - 2])) \<union>
+      path_image (linepath(vector[(pathstart f)$1,a$2 - 2])(pathstart f)) \<union> path_image f \<union>
+      path_image (linepath(pathfinish f)(vector[(pathfinish f)$1,a$2 - 2])) \<union>
+      path_image (linepath(vector[(pathfinish f)$1,a$2 - 2])(vector[b$1 + 2,a$2 - 2]))"
+    "path_image ?P2 = path_image(linepath(vector[(pathstart g)$1, (pathstart g)$2 - 3])(pathstart g)) \<union> path_image g \<union>
+      path_image(linepath(pathfinish g)(vector[(pathfinish g)$1,a$2 - 1])) \<union>
+      path_image(linepath(vector[(pathfinish g)$1,a$2 - 1])(vector[b$1 + 1,a$2 - 1])) \<union>
+      path_image(linepath(vector[b$1 + 1,a$2 - 1])(vector[b$1 + 1,b$2 + 3]))" using assms(1-2)
+      by(auto simp add: path_image_join path_linepath)
+  have abab: "cbox a b \<subseteq> cbox ?a ?b"
+    unfolding interval_cbox_cart[symmetric]
+    by (auto simp add:less_eq_vec_def forall_2 vector_2)
+  obtain z where
+    "z \<in> path_image
+          (linepath (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) +++
+           linepath (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f) +++
+           f +++
+           linepath (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) +++
+           linepath (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]))"
+    "z \<in> path_image
+          (linepath (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g) +++
+           g +++
+           linepath (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1]) +++
+           linepath (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1]) +++
+           linepath (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]))"
+    apply (rule fashoda[of ?P1 ?P2 ?a ?b])
+    unfolding pathstart_join pathfinish_join pathstart_linepath pathfinish_linepath vector_2
+  proof -
+    show "path ?P1" and "path ?P2"
+      using assms by auto
+    have "path_image ?P1 \<subseteq> cbox ?a ?b"
+      unfolding P1P2 path_image_linepath
+      apply (rule Un_least)+
+      defer 3
+      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
+      unfolding mem_interval_cart forall_2 vector_2
+      using ab startfin abab assms(3)
+      using assms(9-)
+      unfolding assms
+      apply (auto simp add: field_simps box_def)
+      done
+    then show "path_image ?P1 \<subseteq> cbox ?a ?b" .
+    have "path_image ?P2 \<subseteq> cbox ?a ?b"
+      unfolding P1P2 path_image_linepath
+      apply (rule Un_least)+
+      defer 2
+      apply (rule_tac[1-4] convex_box(1)[unfolded convex_contains_segment,rule_format])
+      unfolding mem_interval_cart forall_2 vector_2
+      using ab startfin abab assms(4)
+      using assms(9-)
+      unfolding assms
+      apply (auto simp add: field_simps box_def)
+      done
+    then show "path_image ?P2 \<subseteq> cbox ?a ?b" .
+    show "a $ 1 - 2 = a $ 1 - 2"
+      and "b $ 1 + 2 = b $ 1 + 2"
+      and "pathstart g $ 2 - 3 = a $ 2 - 3"
+      and "b $ 2 + 3 = b $ 2 + 3"
+      by (auto simp add: assms)
+  qed
+  note z=this[unfolded P1P2 path_image_linepath]
+  show thesis
+    apply (rule that[of z])
+  proof -
+    have "(z \<in> closed_segment (vector [a $ 1 - 2, a $ 2 - 2]) (vector [pathstart f $ 1, a $ 2 - 2]) \<or>
+      z \<in> closed_segment (vector [pathstart f $ 1, a $ 2 - 2]) (pathstart f)) \<or>
+      z \<in> closed_segment (pathfinish f) (vector [pathfinish f $ 1, a $ 2 - 2]) \<or>
+      z \<in> closed_segment (vector [pathfinish f $ 1, a $ 2 - 2]) (vector [b $ 1 + 2, a $ 2 - 2]) \<Longrightarrow>
+    (((z \<in> closed_segment (vector [pathstart g $ 1, pathstart g $ 2 - 3]) (pathstart g)) \<or>
+      z \<in> closed_segment (pathfinish g) (vector [pathfinish g $ 1, a $ 2 - 1])) \<or>
+      z \<in> closed_segment (vector [pathfinish g $ 1, a $ 2 - 1]) (vector [b $ 1 + 1, a $ 2 - 1])) \<or>
+      z \<in> closed_segment (vector [b $ 1 + 1, a $ 2 - 1]) (vector [b $ 1 + 1, b $ 2 + 3]) \<Longrightarrow> False"
+    proof (simp only: segment_vertical segment_horizontal vector_2, goal_cases)
+      case prems: 1
+      have "pathfinish f \<in> cbox a b"
+        using assms(3) pathfinish_in_path_image[of f] by auto
+      then have "1 + b $ 1 \<le> pathfinish f $ 1 \<Longrightarrow> False"
+        unfolding mem_interval_cart forall_2 by auto
+      then have "z$1 \<noteq> pathfinish f$1"
+        using prems(2)
+        using assms ab
+        by (auto simp add: field_simps)
+      moreover have "pathstart f \<in> cbox a b"
+        using assms(3) pathstart_in_path_image[of f]
+        by auto
+      then have "1 + b $ 1 \<le> pathstart f $ 1 \<Longrightarrow> False"
+        unfolding mem_interval_cart forall_2
+        by auto
+      then have "z$1 \<noteq> pathstart f$1"
+        using prems(2) using assms ab
+        by (auto simp add: field_simps)
+      ultimately have *: "z$2 = a$2 - 2"
+        using prems(1)
+        by auto
+      have "z$1 \<noteq> pathfinish g$1"
+        using prems(2)
+        using assms ab
+        by (auto simp add: field_simps *)
+      moreover have "pathstart g \<in> cbox a b"
+        using assms(4) pathstart_in_path_image[of g]
+        by auto
+      note this[unfolded mem_interval_cart forall_2]
+      then have "z$1 \<noteq> pathstart g$1"
+        using prems(1)
+        using assms ab
+        by (auto simp add: field_simps *)
+      ultimately have "a $ 2 - 1 \<le> z $ 2 \<and> z $ 2 \<le> b $ 2 + 3 \<or> b $ 2 + 3 \<le> z $ 2 \<and> z $ 2 \<le> a $ 2 - 1"
+        using prems(2)
+        unfolding * assms
+        by (auto simp add: field_simps)
+      then show False
+        unfolding * using ab by auto
+    qed
+    then have "z \<in> path_image f \<or> z \<in> path_image g"
+      using z unfolding Un_iff by blast
+    then have z': "z \<in> cbox a b"
+      using assms(3-4)
+      by auto
+    have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart f $ 1 \<or> z $ 1 = pathfinish f $ 1) \<Longrightarrow>
+      z = pathstart f \<or> z = pathfinish f"
+      unfolding vec_eq_iff forall_2 assms
+      by auto
+    with z' show "z \<in> path_image f"
+      using z(1)
+      unfolding Un_iff mem_interval_cart forall_2
+      apply -
+      apply (simp only: segment_vertical segment_horizontal vector_2)
+      unfolding assms
+      apply auto
+      done
+    have "a $ 2 = z $ 2 \<Longrightarrow> (z $ 1 = pathstart g $ 1 \<or> z $ 1 = pathfinish g $ 1) \<Longrightarrow>
+      z = pathstart g \<or> z = pathfinish g"
+      unfolding vec_eq_iff forall_2 assms
+      by auto
+    with z' show "z \<in> path_image g"
+      using z(2)
+      unfolding Un_iff mem_interval_cart forall_2
+      apply (simp only: segment_vertical segment_horizontal vector_2)
+      unfolding assms
+      apply auto
+      done
+  qed
+qed
+
+(** The Following still needs to be translated. Maybe I will do that later.
+
+(* ------------------------------------------------------------------------- *)
+(* Complement in dimension N >= 2 of set homeomorphic to any interval in     *)
+(* any dimension is (path-)connected. This naively generalizes the argument  *)
+(* in Ryuji Maehara's paper "The Jordan curve theorem via the Brouwer        *)
+(* fixed point theorem", American Mathematical Monthly 1984.                 *)
+(* ------------------------------------------------------------------------- *)
+
+let RETRACTION_INJECTIVE_IMAGE_INTERVAL = prove
+ (`!p:real^M->real^N a b.
+        ~(interval[a,b] = {}) /\
+        p continuous_on interval[a,b] /\
+        (!x y. x IN interval[a,b] /\ y IN interval[a,b] /\ p x = p y ==> x = y)
+        ==> ?f. f continuous_on (:real^N) /\
+                IMAGE f (:real^N) SUBSET (IMAGE p (interval[a,b])) /\
+                (!x. x IN (IMAGE p (interval[a,b])) ==> f x = x)`,
+  REPEAT STRIP_TAC THEN
+  FIRST_ASSUM(MP_TAC o GEN_REWRITE_RULE I [INJECTIVE_ON_LEFT_INVERSE]) THEN
+  DISCH_THEN(X_CHOOSE_TAC `q:real^N->real^M`) THEN
+  SUBGOAL_THEN `(q:real^N->real^M) continuous_on
+                (IMAGE p (interval[a:real^M,b]))`
+  ASSUME_TAC THENL
+   [MATCH_MP_TAC CONTINUOUS_ON_INVERSE THEN ASM_REWRITE_TAC[COMPACT_INTERVAL];
+    ALL_TAC] THEN
+  MP_TAC(ISPECL [`q:real^N->real^M`;
+                 `IMAGE (p:real^M->real^N)
+                 (interval[a,b])`;
+                 `a:real^M`; `b:real^M`]
+        TIETZE_CLOSED_INTERVAL) THEN
+  ASM_SIMP_TAC[COMPACT_INTERVAL; COMPACT_CONTINUOUS_IMAGE;
+               COMPACT_IMP_CLOSED] THEN
+  ANTS_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^M` STRIP_ASSUME_TAC) THEN
+  EXISTS_TAC `(p:real^M->real^N) o (r:real^N->real^M)` THEN
+  REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; o_THM; IN_UNIV] THEN
+  CONJ_TAC THENL [ALL_TAC; ASM SET_TAC[]] THEN
+  MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN ASM_REWRITE_TAC[] THEN
+  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP(REWRITE_RULE[IMP_CONJ]
+        CONTINUOUS_ON_SUBSET)) THEN ASM SET_TAC[]);;
+
+let UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
+ (`!s:real^N->bool a b:real^M.
+        s homeomorphic (interval[a,b])
+        ==> !x. ~(x IN s) ==> ~bounded(path_component((:real^N) DIFF s) x)`,
+  REPEAT GEN_TAC THEN REWRITE_TAC[homeomorphic; homeomorphism] THEN
+  REWRITE_TAC[LEFT_IMP_EXISTS_THM] THEN
+  MAP_EVERY X_GEN_TAC [`p':real^N->real^M`; `p:real^M->real^N`] THEN
+  DISCH_TAC THEN
+  SUBGOAL_THEN
+   `!x y. x IN interval[a,b] /\ y IN interval[a,b] /\
+          (p:real^M->real^N) x = p y ==> x = y`
+  ASSUME_TAC THENL [ASM_MESON_TAC[]; ALL_TAC] THEN
+  FIRST_X_ASSUM(MP_TAC o funpow 4 CONJUNCT2) THEN
+  DISCH_THEN(CONJUNCTS_THEN2 (SUBST1_TAC o SYM) ASSUME_TAC) THEN
+  ASM_CASES_TAC `interval[a:real^M,b] = {}` THEN
+  ASM_REWRITE_TAC[IMAGE_CLAUSES; DIFF_EMPTY; PATH_COMPONENT_UNIV;
+                  NOT_BOUNDED_UNIV] THEN
+  ABBREV_TAC `s = (:real^N) DIFF (IMAGE p (interval[a:real^M,b]))` THEN
+  X_GEN_TAC `c:real^N` THEN REPEAT STRIP_TAC THEN
+  SUBGOAL_THEN `(c:real^N) IN s` ASSUME_TAC THENL [ASM SET_TAC[]; ALL_TAC] THEN
+  SUBGOAL_THEN `bounded((path_component s c) UNION
+                        (IMAGE (p:real^M->real^N) (interval[a,b])))`
+  MP_TAC THENL
+   [ASM_SIMP_TAC[BOUNDED_UNION; COMPACT_IMP_BOUNDED; COMPACT_IMP_BOUNDED;
+                 COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+    ALL_TAC] THEN
+  DISCH_THEN(MP_TAC o SPEC `c:real^N` o MATCH_MP BOUNDED_SUBSET_BALL) THEN
+  REWRITE_TAC[UNION_SUBSET] THEN
+  DISCH_THEN(X_CHOOSE_THEN `B:real` STRIP_ASSUME_TAC) THEN
+  MP_TAC(ISPECL [`p:real^M->real^N`; `a:real^M`; `b:real^M`]
+    RETRACTION_INJECTIVE_IMAGE_INTERVAL) THEN
+  ASM_REWRITE_TAC[SUBSET; IN_UNIV] THEN
+  DISCH_THEN(X_CHOOSE_THEN `r:real^N->real^N` MP_TAC) THEN
+  DISCH_THEN(CONJUNCTS_THEN2 ASSUME_TAC
+   (CONJUNCTS_THEN2 MP_TAC ASSUME_TAC)) THEN
+  REWRITE_TAC[FORALL_IN_IMAGE; IN_UNIV] THEN DISCH_TAC THEN
+  ABBREV_TAC `q = \z:real^N. if z IN path_component s c then r(z) else z` THEN
+  SUBGOAL_THEN
+    `(q:real^N->real^N) continuous_on
+     (closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)))`
+  MP_TAC THENL
+   [EXPAND_TAC "q" THEN MATCH_MP_TAC CONTINUOUS_ON_CASES THEN
+    REWRITE_TAC[CLOSED_CLOSURE; CONTINUOUS_ON_ID; GSYM OPEN_CLOSED] THEN
+    REPEAT CONJ_TAC THENL
+     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
+      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
+                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+      ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV];
+      ALL_TAC] THEN
+    X_GEN_TAC `z:real^N` THEN
+    REWRITE_TAC[SET_RULE `~(z IN (s DIFF t) /\ z IN t)`] THEN
+    STRIP_TAC THEN FIRST_X_ASSUM MATCH_MP_TAC THEN
+    MP_TAC(ISPECL
+     [`path_component s (z:real^N)`; `path_component s (c:real^N)`]
+     OPEN_INTER_CLOSURE_EQ_EMPTY) THEN
+    ASM_REWRITE_TAC[GSYM DISJOINT; PATH_COMPONENT_DISJOINT] THEN ANTS_TAC THENL
+     [MATCH_MP_TAC OPEN_PATH_COMPONENT THEN EXPAND_TAC "s" THEN
+      ASM_SIMP_TAC[GSYM CLOSED_OPEN; COMPACT_IMP_CLOSED;
+                   COMPACT_CONTINUOUS_IMAGE; COMPACT_INTERVAL];
+      REWRITE_TAC[DISJOINT; EXTENSION; IN_INTER; NOT_IN_EMPTY] THEN
+      DISCH_THEN(MP_TAC o SPEC `z:real^N`) THEN ASM_REWRITE_TAC[] THEN
+      GEN_REWRITE_TAC (LAND_CONV o ONCE_DEPTH_CONV) [IN] THEN
+      REWRITE_TAC[PATH_COMPONENT_REFL_EQ] THEN ASM SET_TAC[]];
+    ALL_TAC] THEN
+  SUBGOAL_THEN
+   `closure(path_component s c) UNION ((:real^N) DIFF (path_component s c)) =
+    (:real^N)`
+  SUBST1_TAC THENL
+   [MATCH_MP_TAC(SET_RULE `s SUBSET t ==> t UNION (UNIV DIFF s) = UNIV`) THEN
+    REWRITE_TAC[CLOSURE_SUBSET];
+    DISCH_TAC] THEN
+  MP_TAC(ISPECL
+   [`(\x. &2 % c - x) o
+     (\x. c + B / norm(x - c) % (x - c)) o (q:real^N->real^N)`;
+    `cball(c:real^N,B)`]
+    BROUWER) THEN
+  REWRITE_TAC[NOT_IMP; GSYM CONJ_ASSOC; COMPACT_CBALL; CONVEX_CBALL] THEN
+  ASM_SIMP_TAC[CBALL_EQ_EMPTY; REAL_LT_IMP_LE; REAL_NOT_LT] THEN
+  SUBGOAL_THEN `!x. ~((q:real^N->real^N) x = c)` ASSUME_TAC THENL
+   [X_GEN_TAC `x:real^N` THEN EXPAND_TAC "q" THEN
+    REWRITE_TAC[NORM_EQ_0; VECTOR_SUB_EQ] THEN COND_CASES_TAC THEN
+    ASM SET_TAC[PATH_COMPONENT_REFL_EQ];
+    ALL_TAC] THEN
+  REPEAT CONJ_TAC THENL
+   [MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
+    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN CONJ_TAC THENL
+     [ASM_MESON_TAC[CONTINUOUS_ON_SUBSET; SUBSET_UNIV]; ALL_TAC] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_ADD THEN REWRITE_TAC[CONTINUOUS_ON_CONST] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_MUL THEN
+    SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST] THEN
+    REWRITE_TAC[o_DEF; real_div; LIFT_CMUL] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_CMUL THEN
+    MATCH_MP_TAC(REWRITE_RULE[o_DEF] CONTINUOUS_ON_INV) THEN
+    ASM_REWRITE_TAC[FORALL_IN_IMAGE; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+    SUBGOAL_THEN
+     `(\x:real^N. lift(norm(x - c))) = (lift o norm) o (\x. x - c)`
+    SUBST1_TAC THENL [REWRITE_TAC[o_DEF]; ALL_TAC] THEN
+    MATCH_MP_TAC CONTINUOUS_ON_COMPOSE THEN
+    ASM_SIMP_TAC[CONTINUOUS_ON_SUB; CONTINUOUS_ON_ID; CONTINUOUS_ON_CONST;
+                 CONTINUOUS_ON_LIFT_NORM];
+    REWRITE_TAC[SUBSET; FORALL_IN_IMAGE; IN_CBALL; o_THM; dist] THEN
+    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+    REWRITE_TAC[VECTOR_ARITH `c - (&2 % c - (c + x)) = x`] THEN
+    REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
+    ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+    ASM_REAL_ARITH_TAC;
+    REWRITE_TAC[NOT_EXISTS_THM; TAUT `~(c /\ b) <=> c ==> ~b`] THEN
+    REWRITE_TAC[IN_CBALL; o_THM; dist] THEN
+    X_GEN_TAC `x:real^N` THEN DISCH_TAC THEN
+    REWRITE_TAC[VECTOR_ARITH `&2 % c - (c + x') = x <=> --x' = x - c`] THEN
+    ASM_CASES_TAC `(x:real^N) IN path_component s c` THENL
+     [MATCH_MP_TAC(NORM_ARITH `norm(y) < B /\ norm(x) = B ==> ~(--x = y)`) THEN
+      REWRITE_TAC[NORM_MUL; REAL_ABS_DIV; REAL_ABS_NORM] THEN
+      ASM_SIMP_TAC[REAL_DIV_RMUL; NORM_EQ_0; VECTOR_SUB_EQ] THEN
+      ASM_SIMP_TAC[REAL_ARITH `&0 < B ==> abs B = B`] THEN
+      UNDISCH_TAC `path_component s c SUBSET ball(c:real^N,B)` THEN
+      REWRITE_TAC[SUBSET; IN_BALL] THEN ASM_MESON_TAC[dist; NORM_SUB];
+      EXPAND_TAC "q" THEN REWRITE_TAC[] THEN ASM_REWRITE_TAC[] THEN
+      REWRITE_TAC[VECTOR_ARITH `--(c % x) = x <=> (&1 + c) % x = vec 0`] THEN
+      ASM_REWRITE_TAC[DE_MORGAN_THM; VECTOR_SUB_EQ; VECTOR_MUL_EQ_0] THEN
+      SUBGOAL_THEN `~(x:real^N = c)` ASSUME_TAC THENL
+       [ASM_MESON_TAC[PATH_COMPONENT_REFL; IN]; ALL_TAC] THEN
+      ASM_REWRITE_TAC[] THEN
+      MATCH_MP_TAC(REAL_ARITH `&0 < x ==> ~(&1 + x = &0)`) THEN
+      ASM_SIMP_TAC[REAL_LT_DIV; NORM_POS_LT; VECTOR_SUB_EQ]]]);;
+
+let PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL = prove
+ (`!s:real^N->bool a b:real^M.
+        2 <= dimindex(:N) /\ s homeomorphic interval[a,b]
+        ==> path_connected((:real^N) DIFF s)`,
+  REPEAT STRIP_TAC THEN REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
+  FIRST_ASSUM(MP_TAC o MATCH_MP
+    UNBOUNDED_PATH_COMPONENTS_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
+  ASM_REWRITE_TAC[SET_RULE `~(x IN s) <=> x IN (UNIV DIFF s)`] THEN
+  ABBREV_TAC `t = (:real^N) DIFF s` THEN
+  DISCH_TAC THEN MAP_EVERY X_GEN_TAC [`x:real^N`; `y:real^N`] THEN
+  STRIP_TAC THEN FIRST_ASSUM(MP_TAC o MATCH_MP HOMEOMORPHIC_COMPACTNESS) THEN
+  REWRITE_TAC[COMPACT_INTERVAL] THEN
+  DISCH_THEN(MP_TAC o MATCH_MP COMPACT_IMP_BOUNDED) THEN
+  REWRITE_TAC[BOUNDED_POS; LEFT_IMP_EXISTS_THM] THEN
+  X_GEN_TAC `B:real` THEN STRIP_TAC THEN
+  SUBGOAL_THEN `(?u:real^N. u IN path_component t x /\ B < norm(u)) /\
+                (?v:real^N. v IN path_component t y /\ B < norm(v))`
+  STRIP_ASSUME_TAC THENL
+   [ASM_MESON_TAC[BOUNDED_POS; REAL_NOT_LE]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `u:real^N` THEN
+  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_SYM THEN
+  MATCH_MP_TAC PATH_COMPONENT_TRANS THEN EXISTS_TAC `v:real^N` THEN
+  CONJ_TAC THENL [ASM_MESON_TAC[IN]; ALL_TAC] THEN
+  MATCH_MP_TAC PATH_COMPONENT_OF_SUBSET THEN
+  EXISTS_TAC `(:real^N) DIFF cball(vec 0,B)` THEN CONJ_TAC THENL
+   [EXPAND_TAC "t" THEN MATCH_MP_TAC(SET_RULE
+     `s SUBSET t ==> (u DIFF t) SUBSET (u DIFF s)`) THEN
+    ASM_REWRITE_TAC[SUBSET; IN_CBALL_0];
+    MP_TAC(ISPEC `cball(vec 0:real^N,B)`
+       PATH_CONNECTED_COMPLEMENT_BOUNDED_CONVEX) THEN
+    ASM_REWRITE_TAC[BOUNDED_CBALL; CONVEX_CBALL] THEN
+    REWRITE_TAC[PATH_CONNECTED_IFF_PATH_COMPONENT] THEN
+    DISCH_THEN MATCH_MP_TAC THEN
+    ASM_REWRITE_TAC[IN_DIFF; IN_UNIV; IN_CBALL_0; REAL_NOT_LE]]);;
+
+(* ------------------------------------------------------------------------- *)
+(* In particular, apply all these to the special case of an arc.             *)
+(* ------------------------------------------------------------------------- *)
+
+let RETRACTION_ARC = prove
+ (`!p. arc p
+       ==> ?f. f continuous_on (:real^N) /\
+               IMAGE f (:real^N) SUBSET path_image p /\
+               (!x. x IN path_image p ==> f x = x)`,
+  REWRITE_TAC[arc; path; path_image] THEN REPEAT STRIP_TAC THEN
+  MATCH_MP_TAC RETRACTION_INJECTIVE_IMAGE_INTERVAL THEN
+  ASM_REWRITE_TAC[INTERVAL_EQ_EMPTY_CART_1; DROP_VEC; REAL_NOT_LT; REAL_POS]);;
+
+let PATH_CONNECTED_ARC_COMPLEMENT = prove
+ (`!p. 2 <= dimindex(:N) /\ arc p
+       ==> path_connected((:real^N) DIFF path_image p)`,
+  REWRITE_TAC[arc; path] THEN REPEAT STRIP_TAC THEN SIMP_TAC[path_image] THEN
+  MP_TAC(ISPECL [`path_image p:real^N->bool`; `vec 0:real^1`; `vec 1:real^1`]
+    PATH_CONNECTED_COMPLEMENT_HOMEOMORPHIC_INTERVAL) THEN
+  ASM_REWRITE_TAC[path_image] THEN DISCH_THEN MATCH_MP_TAC THEN
+  ONCE_REWRITE_TAC[HOMEOMORPHIC_SYM] THEN
+  MATCH_MP_TAC HOMEOMORPHIC_COMPACT THEN
+  EXISTS_TAC `p:real^1->real^N` THEN ASM_REWRITE_TAC[COMPACT_INTERVAL]);;
+
+let CONNECTED_ARC_COMPLEMENT = prove
+ (`!p. 2 <= dimindex(:N) /\ arc p
+       ==> connected((:real^N) DIFF path_image p)`,
+  SIMP_TAC[PATH_CONNECTED_ARC_COMPLEMENT; PATH_CONNECTED_IMP_CONNECTED]);; *)
+
+end
--- a/src/HOL/Multivariate_Analysis/Gamma.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2969 +0,0 @@
-(*  Title:    HOL/Multivariate_Analysis/Gamma.thy
-    Author:   Manuel Eberl, TU München
-*)
-
-section \<open>The Gamma Function\<close>
-
-theory Gamma
-imports
-  Complex_Transcendental
-  Summation
-  Harmonic_Numbers
-  "~~/src/HOL/Library/Nonpos_Ints"
-  "~~/src/HOL/Library/Periodic_Fun"
-begin
-
-text \<open>
-  Several equivalent definitions of the Gamma function and its
-  most important properties. Also contains the definition and some properties
-  of the log-Gamma function and the Digamma function and the other Polygamma functions.
-
-  Based on the Gamma function, we also prove the Weierstraß product form of the
-  sin function and, based on this, the solution of the Basel problem (the
-  sum over all @{term "1 / (n::nat)^2"}.
-\<close>
-
-lemma pochhammer_eq_0_imp_nonpos_Int:
-  "pochhammer (x::'a::field_char_0) n = 0 \<Longrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
-  by (auto simp: pochhammer_eq_0_iff)
-
-lemma closed_nonpos_Ints [simp]: "closed (\<int>\<^sub>\<le>\<^sub>0 :: 'a :: real_normed_algebra_1 set)"
-proof -
-  have "\<int>\<^sub>\<le>\<^sub>0 = (of_int ` {n. n \<le> 0} :: 'a set)"
-    by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int)
-  also have "closed \<dots>" by (rule closed_of_int_image)
-  finally show ?thesis .
-qed
-
-lemma plus_one_in_nonpos_Ints_imp: "z + 1 \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
-  using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all
-
-lemma of_int_in_nonpos_Ints_iff:
-  "(of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> 0"
-  by (auto simp: nonpos_Ints_def)
-
-lemma one_plus_of_int_in_nonpos_Ints_iff:
-  "(1 + of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> -1"
-proof -
-  have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp
-  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n + 1 \<le> 0" by (subst of_int_in_nonpos_Ints_iff) simp_all
-  also have "\<dots> \<longleftrightarrow> n \<le> -1" by presburger
-  finally show ?thesis .
-qed
-
-lemma one_minus_of_nat_in_nonpos_Ints_iff:
-  "(1 - of_nat n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0"
-proof -
-  have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp
-  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger
-  finally show ?thesis .
-qed
-
-lemma fraction_not_in_ints:
-  assumes "\<not>(n dvd m)" "n \<noteq> 0"
-  shows   "of_int m / of_int n \<notin> (\<int> :: 'a :: {division_ring,ring_char_0} set)"
-proof
-  assume "of_int m / (of_int n :: 'a) \<in> \<int>"
-  then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases)
-  with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps)
-  hence "m = k * n" by (subst (asm) of_int_eq_iff)
-  hence "n dvd m" by simp
-  with assms(1) show False by contradiction
-qed
-
-lemma fraction_not_in_nats:
-  assumes "\<not>n dvd m" "n \<noteq> 0"
-  shows   "of_int m / of_int n \<notin> (\<nat> :: 'a :: {division_ring,ring_char_0} set)"
-proof
-  assume "of_int m / of_int n \<in> (\<nat> :: 'a set)"
-  also note Nats_subset_Ints
-  finally have "of_int m / of_int n \<in> (\<int> :: 'a set)" .
-  moreover have "of_int m / of_int n \<notin> (\<int> :: 'a set)"
-    using assms by (intro fraction_not_in_ints)
-  ultimately show False by contradiction
-qed
-
-lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \<notin> \<int> \<Longrightarrow> z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  by (auto simp: Ints_def nonpos_Ints_def)
-
-lemma double_in_nonpos_Ints_imp:
-  assumes "2 * (z :: 'a :: field_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<or> z + 1/2 \<in> \<int>\<^sub>\<le>\<^sub>0"
-proof-
-  from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases')
-  thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps)
-qed
-
-
-lemma sin_series: "(\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
-proof -
-  from sin_converges[of z] have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z" .
-  also have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z \<longleftrightarrow>
-                 (\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
-    by (subst sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
-       (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE)
-  finally show ?thesis .
-qed
-
-lemma cos_series: "(\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
-proof -
-  from cos_converges[of z] have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z" .
-  also have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z \<longleftrightarrow>
-                 (\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
-    by (subst sums_mono_reindex[of "\<lambda>n. 2*n", symmetric])
-       (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE)
-  finally show ?thesis .
-qed
-
-lemma sin_z_over_z_series:
-  fixes z :: "'a :: {real_normed_field,banach}"
-  assumes "z \<noteq> 0"
-  shows   "(\<lambda>n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)"
-proof -
-  from sin_series[of z] have "(\<lambda>n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z"
-    by (simp add: field_simps scaleR_conv_of_real)
-  from sums_mult[OF this, of "inverse z"] and assms show ?thesis
-    by (simp add: field_simps)
-qed
-
-lemma sin_z_over_z_series':
-  fixes z :: "'a :: {real_normed_field,banach}"
-  assumes "z \<noteq> 0"
-  shows   "(\<lambda>n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)"
-proof -
-  from sums_split_initial_segment[OF sin_converges[of z], of 1]
-    have "(\<lambda>n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp
-  from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps)
-qed
-
-lemma has_field_derivative_sin_z_over_z:
-  fixes A :: "'a :: {real_normed_field,banach} set"
-  shows "((\<lambda>z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)"
-      (is "(?f has_field_derivative ?f') _")
-proof (rule has_field_derivative_at_within)
-  have "((\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n)
-            has_field_derivative (\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)"
-  proof (rule termdiffs_strong)
-    from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1]
-      show "summable (\<lambda>n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def)
-  qed simp
-  also have "(\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f"
-  proof
-    fix z
-    show "(\<Sum>n. of_real (sin_coeff (n+1)) * z^n)  = ?f z"
-      by (cases "z = 0") (insert sin_z_over_z_series'[of z],
-            simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def)
-  qed
-  also have "(\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) =
-                 diffs (\<lambda>n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero)
-  also have "\<dots> = 0" by (simp add: sin_coeff_def diffs_def)
-  finally show "((\<lambda>z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" .
-qed
-
-lemma round_Re_minimises_norm:
-  "norm ((z::complex) - of_int m) \<ge> norm (z - of_int (round (Re z)))"
-proof -
-  let ?n = "round (Re z)"
-  have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)"
-    by (simp add: cmod_def)
-  also have "\<bar>Re z - of_int ?n\<bar> \<le> \<bar>Re z - of_int m\<bar>" by (rule round_diff_minimal)
-  hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \<le> sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)"
-    by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff)
-  also have "\<dots> = norm (z - of_int m)" by (simp add: cmod_def)
-  finally show ?thesis .
-qed
-
-lemma Re_pos_in_ball:
-  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
-  shows   "Re t > 0"
-proof -
-  have "Re (z - t) \<le> norm (z - t)" by (rule complex_Re_le_cmod)
-  also from assms have "\<dots> < Re z / 2" by (simp add: dist_complex_def)
-  finally show "Re t > 0" using assms by simp
-qed
-
-lemma no_nonpos_Int_in_ball_complex:
-  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
-  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases)
-
-lemma no_nonpos_Int_in_ball:
-  assumes "t \<in> ball z (dist z (round (Re z)))"
-  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
-proof
-  assume "t \<in> \<int>\<^sub>\<le>\<^sub>0"
-  then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases)
-  have "dist z (of_int n) \<le> dist z t + dist t (of_int n)" by (rule dist_triangle)
-  also from assms have "dist z t < dist z (round (Re z))" by simp
-  also have "\<dots> \<le> dist z (of_int n)"
-    using round_Re_minimises_norm[of z] by (simp add: dist_complex_def)
-  finally have "dist t (of_int n) > 0" by simp
-  with \<open>t = of_int n\<close> show False by simp
-qed
-
-lemma no_nonpos_Int_in_ball':
-  assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  obtains d where "d > 0" "\<And>t. t \<in> ball z d \<Longrightarrow> t \<notin> \<int>\<^sub>\<le>\<^sub>0"
-proof (rule that)
-  from assms show "setdist {z} \<int>\<^sub>\<le>\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto
-next
-  fix t assume "t \<in> ball z (setdist {z} \<int>\<^sub>\<le>\<^sub>0)"
-  thus "t \<notin> \<int>\<^sub>\<le>\<^sub>0" using setdist_le_dist[of z "{z}" t "\<int>\<^sub>\<le>\<^sub>0"] by force
-qed
-
-lemma no_nonpos_Real_in_ball:
-  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0" and t: "t \<in> ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
-  shows   "t \<notin> \<real>\<^sub>\<le>\<^sub>0"
-using z
-proof (cases "Im z = 0")
-  assume A: "Im z = 0"
-  with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff)
-  with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff)
-next
-  assume A: "Im z \<noteq> 0"
-  have "abs (Im z) - abs (Im t) \<le> abs (Im z - Im t)" by linarith
-  also have "\<dots> = abs (Im (z - t))" by simp
-  also have "\<dots> \<le> norm (z - t)" by (rule abs_Im_le_cmod)
-  also from A t have "\<dots> \<le> abs (Im z) / 2" by (simp add: dist_complex_def)
-  finally have "abs (Im t) > 0" using A by simp
-  thus ?thesis by (force simp add: complex_nonpos_Reals_iff)
-qed
-
-
-subsection \<open>Definitions\<close>
-
-text \<open>
-  We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}.
-  This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its
-  properties more convenient because one does not have to watch out for discontinuities.
-  (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere,
-  whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers)
-
-  We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage
-  that it is a relatively simple limit that converges everywhere. The limit at the poles is 0
-  (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows
-  immediately from the definition.
-\<close>
-
-definition Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)"
-
-definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n"
-
-definition rGamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))"
-
-lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)"
-  and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)"
-  unfolding Gamma_series_def rGamma_series_def by simp_all
-
-lemma rGamma_series_minus_of_nat:
-  "eventually (\<lambda>n. rGamma_series (- of_nat k) n = 0) sequentially"
-  using eventually_ge_at_top[of k]
-  by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff)
-
-lemma Gamma_series_minus_of_nat:
-  "eventually (\<lambda>n. Gamma_series (- of_nat k) n = 0) sequentially"
-  using eventually_ge_at_top[of k]
-  by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff)
-
-lemma Gamma_series'_minus_of_nat:
-  "eventually (\<lambda>n. Gamma_series' (- of_nat k) n = 0) sequentially"
-  using eventually_gt_at_top[of k]
-  by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff)
-
-lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma_series z \<longlonglongrightarrow> 0"
-  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp)
-
-lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series z \<longlonglongrightarrow> 0"
-  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp)
-
-lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series' z \<longlonglongrightarrow> 0"
-  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp)
-
-lemma Gamma_series_Gamma_series':
-  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(\<lambda>n. Gamma_series' z n / Gamma_series z n) \<longlonglongrightarrow> 1"
-proof (rule Lim_transform_eventually)
-  from eventually_gt_at_top[of "0::nat"]
-    show "eventually (\<lambda>n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially"
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n"
-      by (cases n, simp)
-         (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec'
-               dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp)
-    also from n have "\<dots> = z / of_nat n + 1" by (simp add: divide_simps)
-    finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" ..
-  qed
-  have "(\<lambda>x. z / of_nat x) \<longlonglongrightarrow> 0"
-    by (rule tendsto_norm_zero_cancel)
-       (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n],
-        simp add:  norm_divide inverse_eq_divide)
-  from tendsto_add[OF this tendsto_const[of 1]] show "(\<lambda>n. z / of_nat n + 1) \<longlonglongrightarrow> 1" by simp
-qed
-
-
-subsection \<open>Convergence of the Euler series form\<close>
-
-text \<open>
-  We now show that the series that defines the Gamma function in the Euler form converges
-  and that the function defined by it is continuous on the complex halfspace with positive
-  real part.
-
-  We do this by showing that the logarithm of the Euler series is continuous and converges
-  locally uniformly, which means that the log-Gamma function defined by its limit is also
-  continuous.
-
-  This will later allow us to lift holomorphicity and continuity from the log-Gamma
-  function to the inverse of the Gamma function, and from that to the Gamma function itself.
-\<close>
-
-definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))"
-
-definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
-  "ln_Gamma_series' z n =
-     - euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))"
-
-definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> 'a" where
-  "ln_Gamma z = lim (ln_Gamma_series z)"
-
-
-text \<open>
-  We now show that the log-Gamma series converges locally uniformly for all complex numbers except
-  the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this
-  proof:
-  http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm
-\<close>
-
-context
-begin
-
-private lemma ln_Gamma_series_complex_converges_aux:
-  fixes z :: complex and k :: nat
-  assumes z: "z \<noteq> 0" and k: "of_nat k \<ge> 2*norm z" "k \<ge> 2"
-  shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \<le> 2*(norm z + norm z^2) / of_nat k^2"
-proof -
-  let ?k = "of_nat k :: complex" and ?z = "norm z"
-  have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)"
-    by (simp add: algebra_simps)
-  also have "norm ... \<le> ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)"
-    by (subst norm_mult [symmetric], rule norm_triangle_ineq)
-  also have "norm (Ln (1 + -1/?k) - (-1/?k)) \<le> (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))"
-    using k by (intro Ln_approx_linear) (simp add: norm_divide)
-  hence "?z * norm (ln (1-1/?k) + 1/?k) \<le> ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))"
-    by (intro mult_left_mono) simp_all
-  also have "... \<le> (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k
-    by (simp add: field_simps power2_eq_square norm_divide)
-  also have "... \<le> (?z * 2) / of_nat k^2" using k
-    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
-  also have "norm (ln (1+z/?k) - z/?k) \<le> norm (z/?k)^2 / (1 - norm (z/?k))" using k
-    by (intro Ln_approx_linear) (simp add: norm_divide)
-  hence "norm (ln (1+z/?k) - z/?k) \<le> ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)"
-    by (simp add: field_simps norm_divide)
-  also have "... \<le> (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k
-    by (simp add: field_simps power2_eq_square)
-  also have "... \<le> (?z^2 * 2) / of_nat k^2" using k
-    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
-  also note add_divide_distrib [symmetric]
-  finally show ?thesis by (simp only: distrib_left mult.commute)
-qed
-
-lemma ln_Gamma_series_complex_converges:
-  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  assumes d: "d > 0" "\<And>n. n \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> norm (z - of_int n) > d"
-  shows "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n :: complex)"
-proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI')
-  fix e :: real assume e: "e > 0"
-  define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)"
-  define e' where "e' = e / (2*e'')"
-  have "bounded ((\<lambda>t. norm t + norm t^2) ` cball z d)"
-    by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros)
-  hence "bounded ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto
-  hence bdd: "bdd_above ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above)
-
-  with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def
-    by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z])
-  have e'': "norm t + norm t^2 \<le> e''" if "t \<in> ball z d" for t unfolding e''_def using that
-    by (rule cSUP_upper[OF _ bdd])
-  from e z e''_pos have e': "e' > 0" unfolding e'_def
-    by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps)
-
-  have "summable (\<lambda>k. inverse ((real_of_nat k)^2))"
-    by (rule inverse_power_summable) simp
-  from summable_partial_sum_bound[OF this e'] guess M . note M = this
-
-  define N where "N = max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
-  {
-    from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>"
-      by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
-    hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def
-      by (simp_all add: le_of_int_ceiling)
-    also have "... \<le> of_nat N" unfolding N_def
-      by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1)
-    finally have "of_nat N \<ge> 2 * (norm z + d)" .
-    moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all
-    moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n
-      using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def
-      by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps)
-    moreover note calculation
-  } note N = this
-
-  show "\<exists>M. \<forall>t\<in>ball z d. \<forall>m\<ge>M. \<forall>n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e"
-    unfolding dist_complex_def
-  proof (intro exI[of _ N] ballI allI impI)
-    fix t m n assume t: "t \<in> ball z d" and mn: "m \<ge> N" "n > m"
-    from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def)
-    also have "dist z 0 - dist z t \<le> dist 0 t" using dist_triangle[of 0 z t]
-      by (simp add: dist_commute)
-    finally have t_nz: "t \<noteq> 0" by auto
-
-    have "norm t \<le> norm z + norm (t - z)" by (rule norm_triangle_sub)
-    also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute)
-    also have "2 * (norm z + d) \<le> of_nat N" by (rule N)
-    also have "N \<le> m" by (rule mn)
-    finally have norm_t: "2 * norm t < of_nat m" by simp
-
-    have "ln_Gamma_series t m - ln_Gamma_series t n =
-              (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) +
-              ((\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)))"
-      by (simp add: ln_Gamma_series_def algebra_simps)
-    also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) =
-                 (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
-      by (simp add: setsum_diff)
-    also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce
-    also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) =
-                   (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn
-      by (subst setsum_telescope'' [symmetric]) simp_all
-    also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N
-      by (intro setsum_cong_Suc)
-         (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat)
-    also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k
-      using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps)
-    hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) =
-             (\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N
-      by (intro setsum.cong) simp_all
-    also note setsum.distrib [symmetric]
-    also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le>
-      (\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t
-      by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
-    also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
-      by (simp add: setsum_right_distrib)
-    also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
-      by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
-    also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')"
-      by (simp add: e'_def field_simps power2_eq_square)
-    also from e''[OF t] e''_pos e
-      have "\<dots> \<le> e * 1" by (intro mult_left_mono) (simp_all add: field_simps)
-    finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp
-  qed
-qed
-
-end
-
-lemma ln_Gamma_series_complex_converges':
-  assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
-proof -
-  define d' where "d' = Re z"
-  define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)"
-  have "of_int (round d') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 0" using that
-    by (intro nonpos_Ints_of_int) (simp_all add: round_def)
-  with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less)
-
-  have "d < cmod (z - of_int n)" if "n \<in> \<int>\<^sub>\<le>\<^sub>0" for n
-  proof (cases "Re z > 0")
-    case True
-    from nonpos_Ints_nonpos[OF that] have n: "n \<le> 0" by simp
-    from True have "d = Re z/2" by (simp add: d_def d'_def)
-    also from n True have "\<dots> < Re (z - of_int n)" by simp
-    also have "\<dots> \<le> norm (z - of_int n)" by (rule complex_Re_le_cmod)
-    finally show ?thesis .
-  next
-    case False
-    with assms nonpos_Ints_of_int[of "round (Re z)"]
-      have "z \<noteq> of_int (round d')" by (auto simp: not_less)
-    with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def)
-    also have "\<dots> \<le> norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm)
-    finally show ?thesis .
-  qed
-  hence conv: "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
-    by (intro ln_Gamma_series_complex_converges d_pos z) simp_all
-  from d_pos conv show ?thesis by blast
-qed
-
-lemma ln_Gamma_series_complex_converges'': "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> convergent (ln_Gamma_series z)"
-  by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent)
-
-lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z"
-  using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def)
-
-lemma exp_ln_Gamma_series_complex:
-  assumes "n > 0" "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "exp (ln_Gamma_series z n :: complex) = Gamma_series z n"
-proof -
-  from assms obtain m where m: "n = Suc m" by (cases n) blast
-  from assms have "z \<noteq> 0" by (intro notI) auto
-  with assms have "exp (ln_Gamma_series z n) =
-          (of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"
-    unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum)
-  also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)"
-    by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp)
-  also have "... = (\<Prod>k=1..n. z + k) / fact n"
-    by (simp add: fact_setprod)
-    (subst setprod_dividef [symmetric], simp_all add: field_simps)
-  also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n"
-    by (simp add: setprod.atLeast0_atMost_Suc_shift setprod.atLeast_Suc_atMost_Suc_shift)
-  also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)"
-    unfolding pochhammer_setprod
-    by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
-  also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n"
-    unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat)
-  finally show ?thesis .
-qed
-
-
-lemma ln_Gamma_series'_aux:
-  assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(\<lambda>k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums
-              (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s")
-unfolding sums_def
-proof (rule Lim_transform)
-  show "(\<lambda>n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \<longlonglongrightarrow> ?s"
-    (is "?g \<longlonglongrightarrow> _")
-    by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms)
-
-  have A: "eventually (\<lambda>n. (\<Sum>k<n. ?f k) - ?g n = 0) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"
-      by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric],
-          subst atLeastLessThanSuc_atLeastAtMost) simp_all
-    also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
-      by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse)
-    also from n have "\<dots> - ?g n = 0"
-      by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat)
-    finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
-  qed
-  show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
-qed
-
-
-lemma uniformly_summable_deriv_ln_Gamma:
-  assumes z: "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" and d: "d > 0" "d \<le> norm z/2"
-  shows "uniformly_convergent_on (ball z d)
-            (\<lambda>k z. \<Sum>i<k. inverse (of_nat (Suc i)) - inverse (z + of_nat (Suc i)))"
-           (is "uniformly_convergent_on _ (\<lambda>k z. \<Sum>i<k. ?f i z)")
-proof (rule weierstrass_m_test'_ev)
-  {
-    fix t assume t: "t \<in> ball z d"
-    have "norm z = norm (t + (z - t))" by simp
-    have "norm (t + (z - t)) \<le> norm t + norm (z - t)" by (rule norm_triangle_ineq)
-    also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm)
-    finally have A: "norm t > norm z / 2" using z by (simp add: field_simps)
-
-    have "norm t = norm (z + (t - z))" by simp
-    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
-    also from t d have "norm (t - z) \<le> norm z / 2" by (simp add: dist_norm norm_minus_commute)
-    also from z have "\<dots> < norm z" by simp
-    finally have B: "norm t < 2 * norm z" by simp
-    note A B
-  } note ball = this
-
-  show "eventually (\<lambda>n. \<forall>t\<in>ball z d. norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially"
-    using eventually_gt_at_top apply eventually_elim
-  proof safe
-    fix t :: 'a assume t: "t \<in> ball z d"
-    from z ball[OF t] have t_nz: "t \<noteq> 0" by auto
-    fix n :: nat assume n: "n > nat \<lceil>4 * norm z\<rceil>"
-    from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp
-    also from n have "\<dots>  < of_nat n" by linarith
-    finally have n: "of_nat n > 2 * norm t" .
-    hence "of_nat n > norm t" by simp
-    hence t': "t \<noteq> -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc)
-
-    with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))"
-      by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc)
-    also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))"
-      by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc)
-    also {
-      from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)"
-        by (intro divide_left_mono mult_pos_pos) simp_all
-      also have "\<dots> < norm (of_nat (Suc n) / t) - norm (1 :: 'a)"
-        using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc)
-      also have "\<dots> \<le> norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq)
-      finally have "inverse (norm (of_nat (Suc n)/t + 1)) \<le> 4 * norm z / of_nat (Suc n)"
-        using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc)
-    }
-    also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) =
-                 4 * norm z * inverse (of_nat (Suc n)^2)"
-                 by (simp add: divide_simps power2_eq_square del: of_nat_Suc)
-    finally show "norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)"
-      by (simp del: of_nat_Suc)
-  qed
-next
-  show "summable (\<lambda>n. 4 * norm z * inverse ((of_nat (Suc n))^2))"
-    by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable)
-qed
-
-lemma summable_deriv_ln_Gamma:
-  "z \<noteq> (0 :: 'a :: {real_normed_field,banach}) \<Longrightarrow>
-     summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))"
-  unfolding summable_iff_convergent
-  by (rule uniformly_convergent_imp_convergent,
-      rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all
-
-
-definition Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
-  "Polygamma n z = (if n = 0 then
-      (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else
-      (-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))"
-
-abbreviation Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
-  "Digamma \<equiv> Polygamma 0"
-
-lemma Digamma_def:
-  "Digamma z = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni"
-  by (simp add: Polygamma_def)
-
-
-lemma summable_Digamma:
-  assumes "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0"
-  shows   "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
-proof -
-  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
-                       (0 - inverse (z + of_nat 0))"
-    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
-              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
-  from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]]
-    show "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp
-qed
-
-lemma summable_offset:
-  assumes "summable (\<lambda>n. f (n + k) :: 'a :: real_normed_vector)"
-  shows   "summable f"
-proof -
-  from assms have "convergent (\<lambda>m. \<Sum>n<m. f (n + k))" by (simp add: summable_iff_convergent)
-  hence "convergent (\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)))"
-    by (intro convergent_add convergent_const)
-  also have "(\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k))) = (\<lambda>m. \<Sum>n<m+k. f n)"
-  proof
-    fix m :: nat
-    have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto
-    also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)"
-      by (rule setsum.union_disjoint) auto
-    also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
-      by (intro setsum.reindex_cong[of "\<lambda>n. n + k"])
-         (simp, subst image_add_atLeastLessThan, auto)
-    finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan)
-  qed
-  finally have "(\<lambda>a. setsum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. setsum f {..<m + k})"
-    by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset)
-  thus ?thesis by (auto simp: summable_iff_convergent convergent_def)
-qed
-
-lemma Polygamma_converges:
-  fixes z :: "'a :: {real_normed_field,banach}"
-  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
-  shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))"
-proof (rule weierstrass_m_test'_ev)
-  define e where "e = (1 + d / norm z)"
-  define m where "m = nat \<lceil>norm z * e\<rceil>"
-  {
-    fix t assume t: "t \<in> ball z d"
-    have "norm t = norm (z + (t - z))" by simp
-    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
-    also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute)
-    finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def)
-  } note ball = this
-
-  show "eventually (\<lambda>k. \<forall>t\<in>ball z d. norm (inverse ((t + of_nat k)^n)) \<le>
-            inverse (of_nat (k - m)^n)) sequentially"
-    using eventually_gt_at_top[of m] apply eventually_elim
-  proof (intro ballI)
-    fix k :: nat and t :: 'a assume k: "k > m" and t: "t \<in> ball z d"
-    from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff)
-    also have "\<dots> \<le> norm (of_nat k :: 'a) - norm z * e"
-      unfolding m_def by (subst norm_of_nat) linarith
-    also from ball[OF t] have "\<dots> \<le> norm (of_nat k :: 'a) - norm t" by simp
-    also have "\<dots> \<le> norm (of_nat k + t)" by (rule norm_diff_ineq)
-    finally have "inverse ((norm (t + of_nat k))^n) \<le> inverse (real_of_nat (k - m)^n)" using k n
-      by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc)
-    thus "norm (inverse ((t + of_nat k)^n)) \<le> inverse (of_nat (k - m)^n)"
-      by (simp add: norm_inverse norm_power power_inverse)
-  qed
-
-  have "summable (\<lambda>k. inverse ((real_of_nat k)^n))"
-    using inverse_power_summable[of n] n by simp
-  hence "summable (\<lambda>k. inverse ((real_of_nat (k + m - m))^n))" by simp
-  thus "summable (\<lambda>k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset)
-qed
-
-lemma Polygamma_converges':
-  fixes z :: "'a :: {real_normed_field,banach}"
-  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
-  shows "summable (\<lambda>k. inverse ((z + of_nat k)^n))"
-  using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z]
-  by (simp add: summable_iff_convergent)
-
-lemma has_field_derivative_ln_Gamma_complex [derivative_intros]:
-  fixes z :: complex
-  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
-  shows   "(ln_Gamma has_field_derivative Digamma z) (at z)"
-proof -
-  have not_nonpos_Int [simp]: "t \<notin> \<int>\<^sub>\<le>\<^sub>0" if "Re t > 0" for t
-    using that by (auto elim!: nonpos_Ints_cases')
-  from z have z': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" and z'': "z \<noteq> 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I
-     by blast+
-  let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))"
-  let ?f = "\<lambda>z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\<lambda>z. \<Sum>n. ?f' z n"
-  define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
-  from z have d: "d > 0" "norm z/2 \<ge> d" by (auto simp add: complex_nonpos_Reals_iff d_def)
-  have ball: "Im t = 0 \<longrightarrow> Re t > 0" if "dist z t < d" for t
-    using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff)
-  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
-                       (0 - inverse (z + of_nat 0))"
-    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
-              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
-
-  have "((\<lambda>z. \<Sum>n. ?f z n) has_field_derivative ?F' z) (at z)"
-    using d z ln_Gamma_series'_aux[OF z']
-    apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma)
-    apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball
-             simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff
-             simp del: of_nat_Suc)
-    apply (auto simp add: complex_nonpos_Reals_iff)
-    done
-  with z have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative
-                   ?F' z - euler_mascheroni - inverse z) (at z)"
-    by (force intro!: derivative_eq_intros simp: Digamma_def)
-  also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp
-  also from sums have "-inverse z = (\<Sum>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))"
-    by (simp add: sums_iff)
-  also from sums summable_deriv_ln_Gamma[OF z'']
-    have "?F' z + \<dots> =  (\<Sum>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
-    by (subst suminf_add) (simp_all add: add_ac sums_iff)
-  also have "\<dots> - euler_mascheroni = Digamma z" by (simp add: Digamma_def)
-  finally have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z)
-                    has_field_derivative Digamma z) (at z)" .
-  moreover from eventually_nhds_ball[OF d(1), of z]
-    have "eventually (\<lambda>z. ln_Gamma z = (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)"
-  proof eventually_elim
-    fix t assume "t \<in> ball z d"
-    hence "t \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases)
-    from ln_Gamma_series'_aux[OF this]
-      show "ln_Gamma t = (\<Sum>k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff)
-  qed
-  ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
-qed
-
-declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros]
-
-
-lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni"
-  by (simp add: Digamma_def)
-
-lemma Digamma_plus1:
-  assumes "z \<noteq> 0"
-  shows   "Digamma (z+1) = Digamma z + 1/z"
-proof -
-  have sums: "(\<lambda>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))
-                  sums (inverse (z + of_nat 0) - 0)"
-    by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]]
-              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
-  have "Digamma (z+1) = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) -
-          euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac)
-  also have "suminf ?f = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) +
-                         (\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))"
-    using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff)
-  also have "(\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z"
-    using sums by (simp add: sums_iff inverse_eq_divide)
-  finally show ?thesis by (simp add: Digamma_def[of z])
-qed
-
-lemma Polygamma_plus1:
-  assumes "z \<noteq> 0"
-  shows   "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)"
-proof (cases "n = 0")
-  assume n: "n \<noteq> 0"
-  let ?f = "\<lambda>k. inverse ((z + of_nat k) ^ Suc n)"
-  have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\<Sum>k. ?f (k+1))"
-    using n by (simp add: Polygamma_def add_ac)
-  also have "(\<Sum>k. ?f (k+1)) + (\<Sum>k<1. ?f k) = (\<Sum>k. ?f k)"
-    using Polygamma_converges'[OF assms, of "Suc n"] n
-    by (subst suminf_split_initial_segment [symmetric]) simp_all
-  hence "(\<Sum>k. ?f (k+1)) = (\<Sum>k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps)
-  also have "(-1) ^ Suc n * fact n * ((\<Sum>k. ?f k) - inverse (z ^ Suc n)) =
-               Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n
-    by (simp add: inverse_eq_divide algebra_simps Polygamma_def)
-  finally show ?thesis .
-qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide)
-
-lemma Digamma_of_nat:
-  "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni"
-proof (induction n)
-  case (Suc n)
-  have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp
-  also have "\<dots> = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))"
-    by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc)
-  also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc)
-  also have "\<dots> + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni"
-    by (simp add: harm_Suc)
-  finally show ?case .
-qed (simp add: harm_def)
-
-lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni"
-  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl)
-
-lemma Polygamma_of_real: "x \<noteq> 0 \<Longrightarrow> Polygamma n (of_real x) = of_real (Polygamma n x)"
-  unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"]
-  by (simp_all add: suminf_of_real)
-
-lemma Polygamma_Real: "z \<in> \<real> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Polygamma n z \<in> \<real>"
-  by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all
-
-lemma Digamma_half_integer:
-  "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) =
-      (\<Sum>k<n. 2 / (of_nat (2*k+1))) - euler_mascheroni - of_real (2 * ln 2)"
-proof (induction n)
-  case 0
-  have "Digamma (1/2 :: 'a) = of_real (Digamma (1/2))" by (simp add: Polygamma_of_real [symmetric])
-  also have "Digamma (1/2::real) =
-               (\<Sum>k. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni"
-    by (simp add: Digamma_def add_ac)
-  also have "(\<Sum>k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) =
-             (\<Sum>k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))"
-    by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide)
-  also have "\<dots> = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums']
-    by (subst suminf_mult) (simp_all add: algebra_simps sums_iff)
-  finally show ?case by simp
-next
-  case (Suc n)
-  have nz: "2 * of_nat n + (1:: 'a) \<noteq> 0"
-     using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac)
-  hence nz': "of_nat n + (1/2::'a) \<noteq> 0" by (simp add: field_simps)
-  have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp
-  also from nz' have "\<dots> = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)"
-    by (rule Digamma_plus1)
-  also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)"
-    by (subst divide_eq_eq) simp_all
-  also note Suc
-  finally show ?case by (simp add: add_ac)
-qed
-
-lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)"
-  using Digamma_half_integer[of 0] by simp
-
-lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0"
-proof -
-  have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp
-  also have "\<dots> = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp
-  also note euler_mascheroni_less_13_over_22
-  also note ln2_le_25_over_36
-  finally show ?thesis by simp
-qed
-
-
-lemma has_field_derivative_Polygamma [derivative_intros]:
-  fixes z :: "'a :: {real_normed_field,euclidean_space}"
-  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)"
-proof (rule has_field_derivative_at_within, cases "n = 0")
-  assume n: "n = 0"
-  let ?f = "\<lambda>k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)"
-  let ?F = "\<lambda>z. \<Sum>k. ?f k z" and ?f' = "\<lambda>k z. inverse ((z + of_nat k)\<^sup>2)"
-  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
-  from z have summable: "summable (\<lambda>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))"
-    by (intro summable_Digamma) force
-  from z have conv: "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)\<^sup>2))"
-    by (intro Polygamma_converges) auto
-  with d have "summable (\<lambda>k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent
-    by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent )
-
-  have "(?F has_field_derivative (\<Sum>k. ?f' k z)) (at z)"
-  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
-    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
-    from t d(2)[of t] show "((\<lambda>z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)"
-      by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc
-               dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases)
-  qed (insert d(1) summable conv, (assumption|simp)+)
-  with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
-    unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n
-    by (force simp: power2_eq_square intro!: derivative_eq_intros)
-next
-  assume n: "n \<noteq> 0"
-  from z have z': "z \<noteq> 0" by auto
-  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
-  define n' where "n' = Suc n"
-  from n have n': "n' \<ge> 2" by (simp add: n'_def)
-  have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
-                (\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)"
-  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
-    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
-    with d have t': "t \<notin> \<int>\<^sub>\<le>\<^sub>0" "t \<noteq> 0" by auto
-    show "((\<lambda>a. inverse ((a + of_nat k) ^ n')) has_field_derivative
-                - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t'
-      by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp)
-  next
-    have "uniformly_convergent_on (ball z d)
-              (\<lambda>k z. (- of_nat n' :: 'a) * (\<Sum>i<k. inverse ((z + of_nat i) ^ (n'+1))))"
-      using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
-    thus "uniformly_convergent_on (ball z d)
-              (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
-      by (subst (asm) setsum_right_distrib) simp
-  qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
-  also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
-               (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
-    using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all
-  finally have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
-                    - of_nat n' * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" .
-  from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"]
-    show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
-    unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps)
-qed
-
-declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros]
-
-lemma isCont_Polygamma [continuous_intros]:
-  fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
-  shows "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Polygamma n (f x)) z"
-  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Polygamma]])
-
-lemma continuous_on_Polygamma:
-  "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A (Polygamma n :: _ \<Rightarrow> 'a :: {real_normed_field,euclidean_space})"
-  by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast
-
-lemma isCont_ln_Gamma_complex [continuous_intros]:
-  fixes f :: "'a::t2_space \<Rightarrow> complex"
-  shows "isCont f z \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>z. ln_Gamma (f z)) z"
-  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]])
-
-lemma continuous_on_ln_Gamma_complex [continuous_intros]:
-  fixes A :: "complex set"
-  shows "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A ln_Gamma"
-  by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident])
-     fastforce
-
-text \<open>
-  We define a type class that captures all the fundamental properties of the inverse of the Gamma function
-  and defines the Gamma function upon that. This allows us to instantiate the type class both for
-  the reals and for the complex numbers with a minimal amount of proof duplication.
-\<close>
-
-class Gamma = real_normed_field + complete_space +
-  fixes rGamma :: "'a \<Rightarrow> 'a"
-  assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
-  assumes differentiable_rGamma_aux1:
-    "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
-     let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
-               \<longlonglongrightarrow> d) - scaleR euler_mascheroni 1
-     in  filterlim (\<lambda>y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R
-                        norm (y - z)) (nhds 0) (at z)"
-  assumes differentiable_rGamma_aux2:
-    "let z = - of_nat n
-     in  filterlim (\<lambda>y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R
-                        norm (y - z)) (nhds 0) (at z)"
-  assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
-             let fact' = (\<lambda>n. setprod of_nat {1..n});
-                 exp = (\<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x^k /\<^sub>R fact k) \<longlonglongrightarrow> e);
-                 pochhammer' = (\<lambda>a n. (\<Prod>n = 0..n. a + of_nat n))
-             in  filterlim (\<lambda>n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1))))
-                     (nhds (rGamma z)) sequentially"
-begin
-subclass banach ..
-end
-
-definition "Gamma z = inverse (rGamma z)"
-
-
-subsection \<open>Basic properties\<close>
-
-lemma Gamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z = 0"
-  and rGamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z = 0"
-  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
-
-lemma Gamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z \<noteq> 0"
-  and rGamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z \<noteq> 0"
-  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
-
-lemma Gamma_eq_zero_iff: "Gamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
-  and rGamma_eq_zero_iff: "rGamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
-  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
-
-lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)"
-  unfolding Gamma_def by simp
-
-lemma rGamma_series_LIMSEQ [tendsto_intros]:
-  "rGamma_series z \<longlonglongrightarrow> rGamma z"
-proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-  case False
-  hence "z \<noteq> - of_nat n" for n by auto
-  from rGamma_series_aux[OF this] show ?thesis
-    by (simp add: rGamma_series_def[abs_def] fact_setprod pochhammer_Suc_setprod
-                  exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost)
-qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ)
-
-lemma Gamma_series_LIMSEQ [tendsto_intros]:
-  "Gamma_series z \<longlonglongrightarrow> Gamma z"
-proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-  case False
-  hence "(\<lambda>n. inverse (rGamma_series z n)) \<longlonglongrightarrow> inverse (rGamma z)"
-    by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff)
-  also have "(\<lambda>n. inverse (rGamma_series z n)) = Gamma_series z"
-    by (simp add: rGamma_series_def Gamma_series_def[abs_def])
-  finally show ?thesis by (simp add: Gamma_def)
-qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ)
-
-lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)"
-  using Gamma_series_LIMSEQ[of z] by (simp add: limI)
-
-lemma rGamma_1 [simp]: "rGamma 1 = 1"
-proof -
-  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
-                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
-  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
-  moreover have "rGamma_series 1 \<longlonglongrightarrow> rGamma 1" by (rule tendsto_intros)
-  ultimately show ?thesis by (intro LIMSEQ_unique)
-qed
-
-lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z"
-proof -
-  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
-  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
-             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
-      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
-    also from n have "\<dots> = ?f n * rGamma_series z n"
-      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
-    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
-  qed
-  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
-    by (intro tendsto_intros lim_inverse_n)
-  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
-  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
-    by (rule Lim_transform_eventually)
-  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
-    by (intro tendsto_intros)
-  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
-qed
-
-
-lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)"
-proof (induction n arbitrary: z)
-  case (Suc n z)
-  have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH)
-  also note rGamma_plus1 [symmetric]
-  finally show ?case by (simp add: add_ac pochhammer_rec')
-qed simp_all
-
-lemma Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z"
-  using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff)
-
-lemma pochhammer_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> pochhammer z n = Gamma (z + of_nat n) / Gamma z"
-  using pochhammer_rGamma[of z]
-  by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps)
-
-lemma Gamma_0 [simp]: "Gamma 0 = 0"
-  and rGamma_0 [simp]: "rGamma 0 = 0"
-  and Gamma_neg_1 [simp]: "Gamma (- 1) = 0"
-  and rGamma_neg_1 [simp]: "rGamma (- 1) = 0"
-  and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0"
-  and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0"
-  and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0"
-  and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0"
-  by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff)
-
-lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp
-
-lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n"
-  by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff 
-        of_nat_Suc [symmetric] del: of_nat_Suc)
-
-lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)"
-  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, 
-      subst of_nat_Suc, subst Gamma_fact) (rule refl)
-
-lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)"
-proof (cases "n > 0")
-  case True
-  hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all
-  with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp
-qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int)
-
-lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)"
-  by (simp add: Gamma_of_int rGamma_inverse_Gamma)
-
-lemma Gamma_seriesI:
-  assumes "(\<lambda>n. g n / Gamma_series z n) \<longlonglongrightarrow> 1"
-  shows   "g \<longlonglongrightarrow> Gamma z"
-proof (rule Lim_transform_eventually)
-  have "1/2 > (0::real)" by simp
-  from tendstoD[OF assms, OF this]
-    show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially"
-    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
-  from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z"
-    by (intro tendsto_intros)
-  thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp
-qed
-
-lemma Gamma_seriesI':
-  assumes "f \<longlonglongrightarrow> rGamma z"
-  assumes "(\<lambda>n. g n * f n) \<longlonglongrightarrow> 1"
-  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "g \<longlonglongrightarrow> Gamma z"
-proof (rule Lim_transform_eventually)
-  have "1/2 > (0::real)" by simp
-  from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially"
-    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
-  from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z"
-    by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff)
-  thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse)
-qed
-
-lemma Gamma_series'_LIMSEQ: "Gamma_series' z \<longlonglongrightarrow> Gamma z"
-  by (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series']
-                                      Gamma_series'_nonpos_Ints_LIMSEQ[of z])
-
-
-subsection \<open>Differentiability\<close>
-
-lemma has_field_derivative_rGamma_no_nonpos_int:
-  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)"
-proof (rule has_field_derivative_at_within)
-  from assms have "z \<noteq> - of_nat n" for n by auto
-  from differentiable_rGamma_aux1[OF this]
-    show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)"
-         unfolding Digamma_def suminf_def sums_def[abs_def]
-                   has_field_derivative_def has_derivative_def netlimit_at
-    by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric])
-qed
-
-lemma has_field_derivative_rGamma_nonpos_int:
-  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)"
-  apply (rule has_field_derivative_at_within)
-  using differentiable_rGamma_aux2[of n]
-  unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at
-  by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_setprod) simp
-
-lemma has_field_derivative_rGamma [derivative_intros]:
-  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>norm z\<rfloor>) * fact (nat \<lfloor>norm z\<rfloor>)
-      else -rGamma z * Digamma z)) (at z within A)"
-using has_field_derivative_rGamma_no_nonpos_int[of z A]
-      has_field_derivative_rGamma_nonpos_int[of "nat \<lfloor>norm z\<rfloor>" A]
-  by (auto elim!: nonpos_Ints_cases')
-
-declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros]
-declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros]
-declare has_field_derivative_rGamma_nonpos_int [derivative_intros]
-declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros]
-declare has_field_derivative_rGamma [derivative_intros]
-
-
-lemma has_field_derivative_Gamma [derivative_intros]:
-  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)"
-  unfolding Gamma_def [abs_def]
-  by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff)
-
-declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros]
-
-(* TODO: Hide ugly facts properly *)
-hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2
-          differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux
-
-
-
-(* TODO: differentiable etc. *)
-
-
-subsection \<open>Continuity\<close>
-
-lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma"
-  by (rule DERIV_continuous_on has_field_derivative_rGamma)+
-
-lemma continuous_on_Gamma [continuous_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A Gamma"
-  by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast
-
-lemma isCont_rGamma [continuous_intros]:
-  "isCont f z \<Longrightarrow> isCont (\<lambda>x. rGamma (f x)) z"
-  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_rGamma]])
-
-lemma isCont_Gamma [continuous_intros]:
-  "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z"
-  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Gamma]])
-
-
-
-text \<open>The complex Gamma function\<close>
-
-instantiation complex :: Gamma
-begin
-
-definition rGamma_complex :: "complex \<Rightarrow> complex" where
-  "rGamma_complex z = lim (rGamma_series z)"
-
-lemma rGamma_series_complex_converges:
-        "convergent (rGamma_series (z :: complex))" (is "?thesis1")
-  and rGamma_complex_altdef:
-        "rGamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2")
-proof -
-  have "?thesis1 \<and> ?thesis2"
-  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-    case False
-    have "rGamma_series z \<longlonglongrightarrow> exp (- ln_Gamma z)"
-    proof (rule Lim_transform_eventually)
-      from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE)
-      from this(1) uniformly_convergent_imp_convergent[OF this(2), of z]
-        have "ln_Gamma_series z \<longlonglongrightarrow> lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff)
-      thus "(\<lambda>n. exp (-ln_Gamma_series z n)) \<longlonglongrightarrow> exp (- ln_Gamma z)"
-        unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus)
-      from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False
-        show "eventually (\<lambda>n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially"
-        by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def)
-    qed
-    with False show ?thesis
-      by (auto simp: convergent_def rGamma_complex_def intro!: limI)
-  next
-    case True
-    then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases')
-    also have "rGamma_series \<dots> \<longlonglongrightarrow> 0"
-      by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const)
-    finally show ?thesis using True
-      by (auto simp: rGamma_complex_def convergent_def intro!: limI)
-  qed
-  thus "?thesis1" "?thesis2" by blast+
-qed
-
-context
-begin
-
-(* TODO: duplication *)
-private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)"
-proof -
-  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
-  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
-             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
-      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
-    also from n have "\<dots> = ?f n * rGamma_series z n"
-      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
-    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
-  qed
-  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
-    using rGamma_series_complex_converges
-    by (intro tendsto_intros lim_inverse_n)
-       (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def)
-  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
-  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
-    by (rule Lim_transform_eventually)
-  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
-    using rGamma_series_complex_converges
-    by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff)
-  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
-qed
-
-private lemma has_field_derivative_rGamma_complex_no_nonpos_Int:
-  assumes "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
-proof -
-  have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z
-  proof (subst DERIV_cong_ev[OF refl _ refl])
-    from that have "eventually (\<lambda>t. t \<in> ball z (Re z/2)) (nhds z)"
-      by (intro eventually_nhds_in_nhd) simp_all
-    thus "eventually (\<lambda>t. rGamma t = exp (- ln_Gamma t)) (nhds z)"
-      using no_nonpos_Int_in_ball_complex[OF that]
-      by (auto elim!: eventually_mono simp: rGamma_complex_altdef)
-  next
-    have "z \<notin> \<real>\<^sub>\<le>\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff)
-    with that show "((\<lambda>t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)"
-     by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef)
-  qed
-
-  from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
-  proof (induction "nat \<lfloor>1 - Re z\<rfloor>" arbitrary: z)
-    case (Suc n z)
-    from Suc.prems have z: "z \<noteq> 0" by auto
-    from Suc.hyps have "n = nat \<lfloor>- Re z\<rfloor>" by linarith
-    hence A: "n = nat \<lfloor>1 - Re (z + 1)\<rfloor>" by simp
-    from Suc.prems have B: "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp)
-
-    have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1)) z) has_field_derivative
-      -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)"
-      by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps)
-    also have "(\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) = rGamma"
-      by (simp add: rGamma_complex_plus1)
-    also from z have "Digamma (z + 1) * z - 1 = z * Digamma z"
-      by (subst Digamma_plus1) (simp_all add: field_simps)
-    also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z"
-      by (simp add: rGamma_complex_plus1[of z, symmetric])
-    finally show ?case .
-  qed (intro diff, simp)
-qed
-
-private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1"
-proof -
-  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
-                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
-  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
-  thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI)
-qed
-
-private lemma has_field_derivative_rGamma_complex_nonpos_Int:
-  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))"
-proof (induction n)
-  case 0
-  have A: "(0::complex) + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by simp
-  have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)"
-    by (rule derivative_eq_intros DERIV_chain refl
-             has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1)
-    thus ?case by (simp add: rGamma_complex_plus1)
-next
-  case (Suc n)
-  hence A: "(rGamma has_field_derivative (-1)^n * fact n)
-                (at (- of_nat (Suc n) + 1 :: complex))" by simp
-   have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative
-             (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))"
-     by (rule derivative_eq_intros refl A DERIV_chain)+
-        (simp add: algebra_simps rGamma_complex_altdef)
-  thus ?case by (simp add: rGamma_complex_plus1)
-qed
-
-instance proof
-  fix z :: complex show "(rGamma z = 0) \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
-    by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases')
-next
-  fix z :: complex assume "\<And>n. z \<noteq> - of_nat n"
-  hence "z \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases')
-  from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this]
-    show "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
-                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma z +
-              rGamma z * d * (y - z)) /\<^sub>R  cmod (y - z)) \<midarrow>z\<rightarrow> 0"
-      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
-                    netlimit_at of_real_def[symmetric] suminf_def)
-next
-  fix n :: nat
-  from has_field_derivative_rGamma_complex_nonpos_Int[of n]
-  show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} *
-                  (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"
-    by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def)
-next
-  fix z :: complex
-  from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z"
-    by (simp add: convergent_LIMSEQ_iff rGamma_complex_def)
-  thus "let fact' = \<lambda>n. setprod of_nat {1..n};
-            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
-            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
-        in  (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z"
-    by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
-                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
-qed
-
-end
-end
-
-
-lemma Gamma_complex_altdef:
-  "Gamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))"
-  unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus)
-
-lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)"
-proof -
-  have "rGamma_series (cnj z) = (\<lambda>n. cnj (rGamma_series z n))"
-    by (intro ext) (simp_all add: rGamma_series_def exp_cnj)
-  also have "... \<longlonglongrightarrow> cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros)
-  finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI])
-qed
-
-lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)"
-  unfolding Gamma_def by (simp add: cnj_rGamma)
-
-lemma Gamma_complex_real:
-  "z \<in> \<real> \<Longrightarrow> Gamma z \<in> (\<real> :: complex set)" and rGamma_complex_real: "z \<in> \<real> \<Longrightarrow> rGamma z \<in> \<real>"
-  by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma)
-
-lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)"
-  using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast
-
-lemma holomorphic_on_rGamma: "rGamma holomorphic_on A"
-  unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma)
-
-lemma analytic_on_rGamma: "rGamma analytic_on A"
-  unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma)
-
-
-lemma field_differentiable_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma field_differentiable (at z within A)"
-  using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto
-
-lemma holomorphic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A"
-  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma)
-
-lemma analytic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma analytic_on A"
-  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
-     (auto intro!: holomorphic_on_Gamma)
-
-lemma has_field_derivative_rGamma_complex' [derivative_intros]:
-  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-Re z\<rfloor>) * fact (nat \<lfloor>-Re z\<rfloor>) else
-        -rGamma z * Digamma z)) (at z within A)"
-  using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases')
-
-declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros]
-
-
-lemma field_differentiable_Polygamma:
-  fixes z::complex
-  shows
-  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Polygamma n field_differentiable (at z within A)"
-  using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto
-
-lemma holomorphic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n holomorphic_on A"
-  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma)
-
-lemma analytic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n analytic_on A"
-  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
-     (auto intro!: holomorphic_on_Polygamma)
-
-
-
-text \<open>The real Gamma function\<close>
-
-lemma rGamma_series_real:
-  "eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially"
-  using eventually_gt_at_top[of "0 :: nat"]
-proof eventually_elim
-  fix n :: nat assume n: "n > 0"
-  have "Re (rGamma_series (of_real x) n) =
-          Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))"
-    using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real)
-  also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) /
-                              (fact n * (exp (x * ln (real_of_nat n))))))"
-    by (subst exp_of_real) simp
-  also from n have "\<dots> = rGamma_series x n"
-    by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def)
-  finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" ..
-qed
-
-instantiation real :: Gamma
-begin
-
-definition "rGamma_real x = Re (rGamma (of_real x :: complex))"
-
-instance proof
-  fix x :: real
-  have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def)
-  also have "of_real \<dots> = rGamma (of_real x :: complex)"
-    by (intro of_real_Re rGamma_complex_real) simp_all
-  also have "\<dots> = 0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff)
-  also have "\<dots> \<longleftrightarrow> (\<exists>n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases')
-  finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp
-next
-  fix x :: real assume "\<And>n. x \<noteq> - of_nat n"
-  hence x: "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0"
-    by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')
-  then have "x \<noteq> 0" by auto
-  with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
-    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
-                  simp: Polygamma_of_real rGamma_real_def [abs_def])
-  thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k))
-                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma x +
-              rGamma x * d * (y - x)) /\<^sub>R  norm (y - x)) \<midarrow>x\<rightarrow> 0"
-      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
-                    netlimit_at of_real_def[symmetric] suminf_def)
-next
-  fix n :: nat
-  have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))"
-    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
-                  simp: Polygamma_of_real rGamma_real_def [abs_def])
-  thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} *
-                  (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"
-    by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def)
-next
-  fix x :: real
-  have "rGamma_series x \<longlonglongrightarrow> rGamma x"
-  proof (rule Lim_transform_eventually)
-    show "(\<lambda>n. Re (rGamma_series (of_real x) n)) \<longlonglongrightarrow> rGamma x" unfolding rGamma_real_def
-      by (intro tendsto_intros)
-  qed (insert rGamma_series_real, simp add: eq_commute)
-  thus "let fact' = \<lambda>n. setprod of_nat {1..n};
-            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
-            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
-        in  (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x"
-    by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
-                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
-qed
-
-end
-
-
-lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)"
-  unfolding rGamma_real_def using rGamma_complex_real by simp
-
-lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)"
-  unfolding Gamma_def by (simp add: rGamma_complex_of_real)
-
-lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))"
-  by (rule sym, rule limI, rule tendsto_intros)
-
-lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))"
-  by (rule sym, rule limI, rule tendsto_intros)
-
-lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))"
-  using rGamma_complex_real[OF Reals_of_real[of x]]
-  by (elim Reals_cases)
-     (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real)
-
-lemma ln_Gamma_series_complex_of_real:
-  "x > 0 \<Longrightarrow> n > 0 \<Longrightarrow> ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)"
-proof -
-  assume xn: "x > 0" "n > 0"
-  have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k
-    using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps)
-  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real)
-qed
-
-lemma ln_Gamma_real_converges:
-  assumes "(x::real) > 0"
-  shows   "convergent (ln_Gamma_series x)"
-proof -
-  have "(\<lambda>n. ln_Gamma_series (complex_of_real x) n) \<longlonglongrightarrow> ln_Gamma (of_real x)" using assms
-    by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff)
-  moreover from eventually_gt_at_top[of "0::nat"]
-    have "eventually (\<lambda>n. complex_of_real (ln_Gamma_series x n) =
-            ln_Gamma_series (complex_of_real x) n) sequentially"
-    by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms)
-  ultimately have "(\<lambda>n. complex_of_real (ln_Gamma_series x n)) \<longlonglongrightarrow> ln_Gamma (of_real x)"
-    by (subst tendsto_cong) assumption+
-  from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def)
-qed
-
-lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \<Longrightarrow> ln_Gamma_series x \<longlonglongrightarrow> ln_Gamma x"
-  using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff)
-
-lemma ln_Gamma_complex_of_real: "x > 0 \<Longrightarrow> ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)"
-proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually)
-  assume x: "x > 0"
-  show "eventually (\<lambda>n. of_real (ln_Gamma_series x n) =
-            ln_Gamma_series (complex_of_real x) n) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-    by eventually_elim (simp add: ln_Gamma_series_complex_of_real x)
-qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def)
-
-lemma Gamma_real_pos_exp: "x > (0 :: real) \<Longrightarrow> Gamma x = exp (ln_Gamma x)"
-  by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff
-                 ln_Gamma_complex_of_real exp_of_real)
-
-lemma ln_Gamma_real_pos: "x > 0 \<Longrightarrow> ln_Gamma x = ln (Gamma x :: real)"
-  unfolding Gamma_real_pos_exp by simp
-
-lemma Gamma_real_pos: "x > (0::real) \<Longrightarrow> Gamma x > 0"
-  by (simp add: Gamma_real_pos_exp)
-
-lemma has_field_derivative_ln_Gamma_real [derivative_intros]:
-  assumes x: "x > (0::real)"
-  shows "(ln_Gamma has_field_derivative Digamma x) (at x)"
-proof (subst DERIV_cong_ev[OF refl _ refl])
-  from assms show "((Re \<circ> ln_Gamma \<circ> complex_of_real) has_field_derivative Digamma x) (at x)"
-    by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex
-             simp: Polygamma_of_real o_def)
-  from eventually_nhds_in_nhd[of x "{0<..}"] assms
-    show "eventually (\<lambda>y. ln_Gamma y = (Re \<circ> ln_Gamma \<circ> of_real) y) (nhds x)"
-    by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open)
-qed
-
-declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros]
-
-
-lemma has_field_derivative_rGamma_real' [derivative_intros]:
-  "(rGamma has_field_derivative (if x \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-x\<rfloor>) * fact (nat \<lfloor>-x\<rfloor>) else
-        -rGamma x * Digamma x)) (at x within A)"
-  using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases')
-
-declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros]
-
-lemma Polygamma_real_odd_pos:
-  assumes "(x::real) \<notin> \<int>\<^sub>\<le>\<^sub>0" "odd n"
-  shows   "Polygamma n x > 0"
-proof -
-  from assms have "x \<noteq> 0" by auto
-  with assms show ?thesis
-    unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
-    by (auto simp: zero_less_power_eq simp del: power_Suc
-             dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos)
-qed
-
-lemma Polygamma_real_even_neg:
-  assumes "(x::real) > 0" "n > 0" "even n"
-  shows   "Polygamma n x < 0"
-  using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
-  by (auto intro!: mult_pos_pos suminf_pos)
-
-lemma Polygamma_real_strict_mono:
-  assumes "x > 0" "x < (y::real)" "even n"
-  shows   "Polygamma n x < Polygamma n y"
-proof -
-  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
-    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
-  then guess \<xi> by (elim exE conjE) note \<xi> = this
-  note \<xi>(3)
-  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> > 0"
-    by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases)
-  finally show ?thesis by simp
-qed
-
-lemma Polygamma_real_strict_antimono:
-  assumes "x > 0" "x < (y::real)" "odd n"
-  shows   "Polygamma n x > Polygamma n y"
-proof -
-  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
-    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
-  then guess \<xi> by (elim exE conjE) note \<xi> = this
-  note \<xi>(3)
-  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> < 0"
-    by (intro mult_pos_neg Polygamma_real_even_neg) simp_all
-  finally show ?thesis by simp
-qed
-
-lemma Polygamma_real_mono:
-  assumes "x > 0" "x \<le> (y::real)" "even n"
-  shows   "Polygamma n x \<le> Polygamma n y"
-  using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2)
-  by (cases "x = y") simp_all
-
-lemma Digamma_real_ge_three_halves_pos:
-  assumes "x \<ge> 3/2"
-  shows   "Digamma (x :: real) > 0"
-proof -
-  have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos)
-  also from assms have "\<dots> \<le> Digamma x" by (intro Polygamma_real_mono) simp_all
-  finally show ?thesis .
-qed
-
-lemma ln_Gamma_real_strict_mono:
-  assumes "x \<ge> 3/2" "x < y"
-  shows   "ln_Gamma (x :: real) < ln_Gamma y"
-proof -
-  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> ln_Gamma y - ln_Gamma x = (y - x) * Digamma \<xi>"
-    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
-  then guess \<xi> by (elim exE conjE) note \<xi> = this
-  note \<xi>(3)
-  also from \<xi>(1,2) assms have "(y - x) * Digamma \<xi> > 0"
-    by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all
-  finally show ?thesis by simp
-qed
-
-lemma Gamma_real_strict_mono:
-  assumes "x \<ge> 3/2" "x < y"
-  shows   "Gamma (x :: real) < Gamma y"
-proof -
-  from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp
-  also have "\<dots> < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms)
-  also from Gamma_real_pos_exp[of y] assms have "\<dots> = Gamma y" by simp
-  finally show ?thesis .
-qed
-
-lemma log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)"
-  by (rule convex_on_realI[of _ _ Digamma])
-     (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos
-           simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases')
-
-
-subsection \<open>Beta function\<close>
-
-definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)"
-
-lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)"
-  by (simp add: inverse_eq_divide Beta_def Gamma_def)
-
-lemma Beta_commute: "Beta a b = Beta b a"
-  unfolding Beta_def by (simp add: ac_simps)
-
-lemma has_field_derivative_Beta1 [derivative_intros]:
-  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "((\<lambda>x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y))))
-               (at x within A)" unfolding Beta_altdef
-  by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps)
-
-lemma Beta_pole1: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
-  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
-
-lemma Beta_pole2: "y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
-  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
-  
-lemma Beta_zero: "x + y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
-  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
-  
-lemma has_field_derivative_Beta2 [derivative_intros]:
-  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "((\<lambda>y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y))))
-               (at y within A)"
-  using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac)
-
-lemma Beta_plus1_plus1:
-  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "Beta (x + 1) y + Beta x (y + 1) = Beta x y"
-proof -
-  have "Beta (x + 1) y + Beta x (y + 1) =
-            (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)"
-    by (simp add: Beta_altdef add_divide_distrib algebra_simps)
-  also have "\<dots> = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))"
-    by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps)
-  also from assms have "\<dots> = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp
-  finally show ?thesis .
-qed
-
-lemma Beta_plus1_left:
-  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(x + y) * Beta (x + 1) y = x * Beta x y"
-proof -
-  have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))"
-    unfolding Beta_altdef by (simp only: ac_simps)
-  also have "\<dots> = x * Beta x y" unfolding Beta_altdef
-     by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps)
-  finally show ?thesis .
-qed
-
-lemma Beta_plus1_right:
-  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(x + y) * Beta x (y + 1) = y * Beta x y"
-  using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute)
-
-lemma Gamma_Gamma_Beta:
-  assumes "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "Gamma x * Gamma y = Beta x y * Gamma (x + y)"
-  unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"]
-  by (simp add: rGamma_inverse_Gamma)
-
-
-
-subsection \<open>Legendre duplication theorem\<close>
-
-context
-begin
-
-private lemma Gamma_legendre_duplication_aux:
-  fixes z :: "'a :: Gamma"
-  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)"
-proof -
-  let ?powr = "\<lambda>b a. exp (a * of_real (ln (of_nat b)))"
-  let ?h = "\<lambda>n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) *
-                exp (1/2 * of_real (ln (real_of_nat n)))"
-  {
-    fix z :: 'a assume z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-    let ?g = "\<lambda>n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n /
-                      Gamma_series' (2*z) (2*n)"
-    have "eventually (\<lambda>n. ?g n = ?h n) sequentially" using eventually_gt_at_top
-    proof eventually_elim
-      fix n :: nat assume n: "n > 0"
-      let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a"
-      have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp
-      have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) /
-                (pochhammer z n * pochhammer (z + 1/2) n)"
-        by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac)
-      have B: "Gamma_series' (2*z) (2*n) =
-                       ?f' * ?powr 2 (2*z) * ?powr n (2*z) /
-                       (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n
-        by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double)
-      from z have "pochhammer z n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
-      moreover from z have "pochhammer (z + 1/2) n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
-      ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) =
-         ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))"
-        using n unfolding A B by (simp add: divide_simps exp_minus)
-      also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)"
-        by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib)
-      finally show "?g n = ?h n" by (simp only: mult_ac)
-    qed
-
-    moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
-    hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
-      using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"]
-      by (intro tendsto_intros Gamma_series'_LIMSEQ)
-         (simp_all add: o_def subseq_def Gamma_eq_zero_iff)
-    ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
-      by (rule Lim_transform_eventually)
-  } note lim = this
-
-  from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
-  from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-    by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all
-  with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real)
-  from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis
-    by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
-qed
-
-(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is
-   infinitely differentiable *)
-private lemma Gamma_reflection_aux:
-  defines "h \<equiv> \<lambda>z::complex. if z \<in> \<int> then 0 else
-                 (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
-  defines "a \<equiv> complex_of_real pi"
-  obtains h' where "continuous_on UNIV h'" "\<And>z. (h has_field_derivative (h' z)) (at z)"
-proof -
-  define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n
-  define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z
-  define g where "g n = complex_of_real (sin_coeff (n+1))" for n
-  define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z
-  have a_nz: "a \<noteq> 0" unfolding a_def by simp
-
-  have "(\<lambda>n. f n * (a*z)^n) sums (F z) \<and> (\<lambda>n. g n * (a*z)^n) sums (G z)"
-    if "abs (Re z) < 1" for z
-  proof (cases "z = 0"; rule conjI)
-    assume "z \<noteq> 0"
-    note z = this that
-
-    from z have sin_nz: "sin (a*z) \<noteq> 0" unfolding a_def by (auto simp: sin_eq_0)
-    have "(\<lambda>n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"]
-      by (simp add: scaleR_conv_of_real)
-    from sums_split_initial_segment[OF this, of 1]
-      have "(\<lambda>n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac)
-    from sums_mult[OF this, of "inverse (a*z)"] z a_nz
-      have A: "(\<lambda>n. g n * (a*z)^n) sums (sin (a*z)/(a*z))"
-      by (simp add: field_simps g_def)
-    with z show "(\<lambda>n. g n * (a*z)^n) sums (G z)" by (simp add: G_def)
-    from A z a_nz sin_nz have g_nz: "(\<Sum>n. g n * (a*z)^n) \<noteq> 0" by (simp add: sums_iff g_def)
-
-    have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def)
-    from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1]
-    have "(\<lambda>n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))"
-      by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def)
-    from sums_mult[OF this, of "inverse z"] z assms
-      show "(\<lambda>n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def)
-  next
-    assume z: "z = 0"
-    have "(\<lambda>n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp
-    with z show "(\<lambda>n. f n * (a * z) ^ n) sums (F z)"
-      by (simp add: f_def F_def sin_coeff_def cos_coeff_def)
-    have "(\<lambda>n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp
-    with z show "(\<lambda>n. g n * (a * z) ^ n) sums (G z)"
-      by (simp add: g_def G_def sin_coeff_def cos_coeff_def)
-  qed
-  note sums = conjunct1[OF this] conjunct2[OF this]
-
-  define h2 where [abs_def]:
-    "h2 z = (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z
-  define POWSER where [abs_def]: "POWSER f z = (\<Sum>n. f n * (z^n :: complex))" for f z
-  define POWSER' where [abs_def]: "POWSER' f z = (\<Sum>n. diffs f n * (z^n))" for f and z :: complex
-  define h2' where [abs_def]:
-    "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) /
-      (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z
-
-  have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t
-  proof -
-    from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases simp: dist_0_norm)
-    hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)"
-      unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def)
-    also have "a*cot (a*t) - 1/t = (F t) / (G t)"
-      using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def)
-    also have "\<dots> = (\<Sum>n. f n * (a*t)^n) / (\<Sum>n. g n * (a*t)^n)"
-      using sums[of t] that by (simp add: sums_iff dist_0_norm)
-    finally show "h t = h2 t" by (simp only: h2_def)
-  qed
-
-  let ?A = "{z. abs (Re z) < 1}"
-  have "open ({z. Re z < 1} \<inter> {z. Re z > -1})"
-    using open_halfspace_Re_gt open_halfspace_Re_lt by auto
-  also have "({z. Re z < 1} \<inter> {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto
-  finally have open_A: "open ?A" .
-  hence [simp]: "interior ?A = ?A" by (simp add: interior_open)
-
-  have summable_f: "summable (\<lambda>n. f n * z^n)" for z
-    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
-       (simp_all add: norm_mult a_def del: of_real_add)
-  have summable_g: "summable (\<lambda>n. g n * z^n)" for z
-    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
-       (simp_all add: norm_mult a_def del: of_real_add)
-  have summable_fg': "summable (\<lambda>n. diffs f n * z^n)" "summable (\<lambda>n. diffs g n * z^n)" for z
-    by (intro termdiff_converges_all summable_f summable_g)+
-  have "(POWSER f has_field_derivative (POWSER' f z)) (at z)"
-               "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z
-    unfolding POWSER_def POWSER'_def
-    by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+
-  note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def]
-  have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z"
-    for z unfolding POWSER_def POWSER'_def
-    by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+
-  note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def]
-
-  {
-    fix z :: complex assume z: "abs (Re z) < 1"
-    define d where "d = \<i> * of_real (norm z + 1)"
-    have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add)
-    have "eventually (\<lambda>z. h z = h2 z) (nhds z)"
-      using eventually_nhds_in_nhd[of z ?A] using h_eq z
-      by (auto elim!: eventually_mono simp: dist_0_norm)
-
-    moreover from sums(2)[OF z] z have nz: "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
-      unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def)
-    have A: "z \<in> \<int> \<longleftrightarrow> z = 0" using z by (auto elim!: Ints_cases)
-    have no_int: "1 + z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of "1+z" 1] A
-      by (auto elim!: nonpos_Ints_cases)
-    have no_int': "1 - z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of 1 "1-z"] A
-      by (auto elim!: nonpos_Ints_cases)
-    from no_int no_int' have no_int: "1 - z \<notin> \<int>\<^sub>\<le>\<^sub>0" "1 + z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
-    have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def
-      by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+)
-         (auto simp: h2'_def POWSER_def field_simps power2_eq_square)
-    ultimately have deriv: "(h has_field_derivative h2' z) (at z)"
-      by (subst DERIV_cong_ev[OF refl _ refl])
-
-    from sums(2)[OF z] z have "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
-      unfolding G_def by (auto simp: sums_iff a_def sin_eq_0)
-    hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def
-      by (intro continuous_intros cont
-            continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto
-    note deriv and this
-  } note A = this
-
-  interpret h: periodic_fun_simple' h
-  proof
-    fix z :: complex
-    show "h (z + 1) = h z"
-    proof (cases "z \<in> \<int>")
-      assume z: "z \<notin> \<int>"
-      hence A: "z + 1 \<notin> \<int>" "z \<noteq> 0" using Ints_diff[of "z+1" 1] by auto
-      hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)"
-        by (subst (1 2) Digamma_plus1) simp_all
-      with A z show "h (z + 1) = h z"
-        by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def)
-    qed (simp add: h_def)
-  qed
-
-  have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z
-  proof -
-    have "((\<lambda>z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)"
-      by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
-         (insert z, auto intro!: derivative_eq_intros)
-    hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1)
-    moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all
-    ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique)
-  qed
-
-  define h2'' where "h2'' z = h2' (z - of_int \<lfloor>Re z\<rfloor>)" for z
-  have deriv: "(h has_field_derivative h2'' z) (at z)" for z
-  proof -
-    fix z :: complex
-    have B: "\<bar>Re z - real_of_int \<lfloor>Re z\<rfloor>\<bar> < 1" by linarith
-    have "((\<lambda>t. h (t - of_int \<lfloor>Re z\<rfloor>)) has_field_derivative h2'' z) (at z)"
-      unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
-                            (insert B, auto intro!: derivative_intros)
-    thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int)
-  qed
-
-  have cont: "continuous_on UNIV h2''"
-  proof (intro continuous_at_imp_continuous_on ballI)
-    fix z :: complex
-    define r where "r = \<lfloor>Re z\<rfloor>"
-    define A where "A = {t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}"
-    have "continuous_on A (\<lambda>t. h2' (t - of_int r))" unfolding A_def
-      by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros)
-         (simp_all add: abs_real_def)
-    moreover have "h2'' t = h2' (t - of_int r)" if t: "t \<in> A" for t
-    proof (cases "Re t \<ge> of_int r")
-      case True
-      from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def)
-      with True have "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor>" unfolding r_def by linarith
-      thus ?thesis by (auto simp: r_def h2''_def)
-    next
-      case False
-      from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def)
-      with False have t': "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor> - 1" unfolding r_def by linarith
-      moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)"
-        by (intro h2'_eq) simp_all
-      ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t')
-    qed
-    ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl])
-    moreover {
-      have "open ({t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t})"
-        by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt)
-      also have "{t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t} = A"
-        unfolding A_def by blast
-      finally have "open A" .
-    }
-    ultimately have C: "isCont h2'' t" if "t \<in> A" for t using that
-      by (subst (asm) continuous_on_eq_continuous_at) auto
-    have "of_int r - 1 < Re z" "Re z  < of_int r + 1" unfolding r_def by linarith+
-    thus "isCont h2'' z" by (intro C) (simp_all add: A_def)
-  qed
-
-  from that[OF cont deriv] show ?thesis .
-qed
-
-lemma Gamma_reflection_complex:
-  fixes z :: complex
-  shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)"
-proof -
-  let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"
-  define g where [abs_def]: "g z = (if z \<in> \<int> then of_real pi else ?g z)" for z :: complex
-  let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
-  define h where [abs_def]: "h z = (if z \<in> \<int> then 0 else ?h z)" for z :: complex
-
-  \<comment> \<open>@{term g} is periodic with period 1.\<close>
-  interpret g: periodic_fun_simple' g
-  proof
-    fix z :: complex
-    show "g (z + 1) = g z"
-    proof (cases "z \<in> \<int>")
-      case False
-      hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def)
-      also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)"
-        using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints
-        by (subst Beta_plus1_left [symmetric]) auto
-      also have "\<dots> * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))"
-        using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints
-        by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi)
-      also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)"
-        using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def)
-      finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto
-    qed (simp add: g_def)
-  qed
-
-  \<comment> \<open>@{term g} is entire.\<close>
-  have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex
-  proof (cases "z \<in> \<int>")
-    let ?h' = "\<lambda>z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) +
-                     of_real pi * cos (z * of_real pi))"
-    case False
-    from False have "eventually (\<lambda>t. t \<in> UNIV - \<int>) (nhds z)"
-      by (intro eventually_nhds_in_open) (auto simp: open_Diff)
-    hence "eventually (\<lambda>t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def)
-    moreover {
-      from False Ints_diff[of 1 "1-z"] have "1 - z \<notin> \<int>" by auto
-      hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints
-        by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def)
-      also from False have "sin (of_real pi * z) \<noteq> 0" by (subst sin_eq_0) auto
-      hence "?h' z = h z * g z"
-        using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def)
-      finally have "(?g has_field_derivative (h z * g z)) (at z)" .
-    }
-    ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
-  next
-    case True
-    then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases)
-    let ?t = "(\<lambda>z::complex. if z = 0 then 1 else sin z / z) \<circ> (\<lambda>z. of_real pi * z)"
-    have deriv_0: "(g has_field_derivative 0) (at 0)"
-    proof (subst DERIV_cong_ev[OF refl _ refl])
-      show "eventually (\<lambda>z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)"
-        using eventually_nhds_ball[OF zero_less_one, of "0::complex"]
-      proof eventually_elim
-        fix z :: complex assume z: "z \<in> ball 0 1"
-        show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z"
-        proof (cases "z = 0")
-          assume z': "z \<noteq> 0"
-          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm)
-          from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp
-          with z'' z' show ?thesis by (simp add: g_def ac_simps)
-        qed (simp add: g_def)
-      qed
-      have "(?t has_field_derivative (0 * of_real pi)) (at 0)"
-        using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"]
-        by (intro DERIV_chain) simp_all
-      thus "((\<lambda>z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)"
-        by (auto intro!: derivative_eq_intros simp: o_def)
-    qed
-
-    have "((g \<circ> (\<lambda>x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))"
-      using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros)
-    also have "g \<circ> (\<lambda>x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int)
-    finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def)
-  qed
-
-  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z
-  proof (cases "z \<in> \<int>")
-    case True
-    with that have "z = 0 \<or> z = 1" by (force elim!: Ints_cases)
-    moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0"
-      using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square)
-    moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1"
-        using fraction_not_in_ints[where 'a = complex, of 2 1]
-        by (simp add: g_def power2_eq_square Beta_def algebra_simps)
-    ultimately show ?thesis by force
-  next
-    case False
-    hence z: "z/2 \<notin> \<int>" "(z+1)/2 \<notin> \<int>" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases)
-    hence z': "z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "(z+1)/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
-    from z have "1-z/2 \<notin> \<int>" "1-((z+1)/2) \<notin> \<int>"
-      using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto
-    hence z'': "1-z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "1-((z+1)/2) \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
-    from z have "g (z/2) * g ((z+1)/2) =
-      (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) *
-      (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))"
-      by (simp add: g_def)
-    also from z' Gamma_legendre_duplication_aux[of "z/2"]
-      have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z"
-      by (simp add: add_divide_distrib)
-    also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"]
-      have "Gamma (1-z/2) * Gamma (1-(z+1)/2) =
-              Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))"
-      by (simp add: add_divide_distrib ac_simps)
-    finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) *
-                    (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))"
-      by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real)
-    also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)"
-      using cos_sin_eq[of "- of_real pi * z/2", symmetric]
-      by (simp add: ring_distribs add_divide_distrib ac_simps)
-    also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)"
-      by (subst sin_times_cos) (simp add: field_simps)
-    also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z"
-      using \<open>z \<notin> \<int>\<close> by (simp add: g_def)
-    finally show ?thesis .
-  qed
-  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z
-  proof -
-    define r where "r = \<lfloor>Re z / 2\<rfloor>"
-    have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int)
-    also have "of_int (2*r) = 2 * of_int r" by simp
-    also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+
-    hence "Gamma (1/2)^2 * g (z - 2 * of_int r) =
-                   g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)"
-      unfolding r_def by (intro g_eq[symmetric]) simp_all
-    also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp
-    also have "g \<dots> = g (z/2)" by (rule g.minus_of_int)
-    also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp
-    also have "g \<dots> = g ((z+1)/2)" by (rule g.minus_of_int)
-    finally show ?thesis ..
-  qed
-
-  have g_nz [simp]: "g z \<noteq> 0" for z :: complex
-  unfolding g_def using Ints_diff[of 1 "1 - z"]
-    by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int)
-
-  have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z
-  proof -
-    have "((\<lambda>t. g (t/2) * g ((t+1)/2)) has_field_derivative
-                       (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
-      by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps)
-    hence "((\<lambda>t. Gamma (1/2)^2 * g t) has_field_derivative
-              Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
-      by (subst (1 2) g_eq[symmetric]) simp
-    from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"]
-      have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)"
-      using fraction_not_in_ints[where 'a = complex, of 2 1]
-      by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints)
-    moreover have "(g has_field_derivative (g z * h z)) (at z)"
-      using g_g'[of z] by (simp add: ac_simps)
-    ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)"
-      by (intro DERIV_unique)
-    thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp
-  qed
-
-  obtain h' where h'_cont: "continuous_on UNIV h'" and
-                  h_h': "\<And>z. (h has_field_derivative h' z) (at z)"
-     unfolding h_def by (erule Gamma_reflection_aux)
-
-  have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z
-  proof -
-    have "((\<lambda>t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative
-                       ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)"
-      by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2])
-    hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)"
-      by (subst (asm) h_eq[symmetric])
-    from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique)
-  qed
-
-  have h'_zero: "h' z = 0" for z
-  proof -
-    define m where "m = max 1 \<bar>Re z\<bar>"
-    define B where "B = {t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
-    have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter>
-                  {t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})"
-      (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le
-                                 closed_halfspace_Im_ge closed_halfspace_Im_le)
-    also have "?B = B" unfolding B_def by fastforce
-    finally have "closed B" .
-    moreover have "bounded B" unfolding bounded_iff
-    proof (intro ballI exI)
-      fix t assume t: "t \<in> B"
-      have "norm t \<le> \<bar>Re t\<bar> + \<bar>Im t\<bar>" by (rule cmod_le)
-      also from t have "\<bar>Re t\<bar> \<le> m" unfolding B_def by blast
-      also from t have "\<bar>Im t\<bar> \<le> \<bar>Im z\<bar>" unfolding B_def by blast
-      finally show "norm t \<le> m + \<bar>Im z\<bar>" by - simp
-    qed
-    ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast
-
-    define M where "M = (SUP z:B. norm (h' z))"
-    have "compact (h' ` B)"
-      by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+
-    hence bdd: "bdd_above ((\<lambda>z. norm (h' z)) ` B)"
-      using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded)
-    have "norm (h' z) \<le> M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def)
-    also have "M \<le> M/2"
-    proof (subst M_def, subst cSUP_le_iff)
-      have "z \<in> B" unfolding B_def m_def by simp
-      thus "B \<noteq> {}" by auto
-    next
-      show "\<forall>z\<in>B. norm (h' z) \<le> M/2"
-      proof
-        fix t :: complex assume t: "t \<in> B"
-        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm)
-        also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp
-        also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))"
-          by (rule norm_triangle_ineq)
-        also from t have "abs (Re ((t + 1)/2)) \<le> m" unfolding m_def B_def by auto
-        with t have "t/2 \<in> B" "(t+1)/2 \<in> B" unfolding B_def by auto
-        hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \<le> M + M" unfolding M_def
-          by (intro add_mono cSUP_upper bdd) (auto simp: B_def)
-        also have "(M + M) / 4 = M / 2" by simp
-        finally show "norm (h' t) \<le> M/2" by - simp_all
-      qed
-    qed (insert bdd, auto simp: cball_eq_empty)
-    hence "M \<le> 0" by simp
-    finally show "h' z = 0" by simp
-  qed
-  have h_h'_2: "(h has_field_derivative 0) (at z)" for z
-    using h_h'[of z] h'_zero[of z] by simp
-
-  have g_real: "g z \<in> \<real>" if "z \<in> \<real>" for z
-    unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real)
-  have h_real: "h z \<in> \<real>" if "z \<in> \<real>" for z
-    unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real)
-  have g_nz: "g z \<noteq> 0" for z unfolding g_def using Ints_diff[of 1 "1-z"]
-    by (auto simp: Gamma_eq_zero_iff sin_eq_0)
-
-  from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c"
-    by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm)
-  then obtain c where c: "\<And>z. h z = c" by auto
-  have "\<exists>u. u \<in> closed_segment 0 1 \<and> Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))"
-    by (intro complex_mvt_line g_g')
-    find_theorems name:deriv Reals
-  then guess u by (elim exE conjE) note u = this
-  from u(1) have u': "u \<in> \<real>" unfolding closed_segment_def
-    by (auto simp: scaleR_conv_of_real)
-  from u' g_real[of u] g_nz[of u] have "Re (g u) \<noteq> 0" by (auto elim!: Reals_cases)
-  with u(2) c[of u] g_real[of u] g_nz[of u] u'
-    have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1)
-  with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases)
-  with c have A: "h z * g z = 0" for z by simp
-  hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
-  hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
-  then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
-  from this[of 0] have "c' = pi" unfolding g_def by simp
-  with c have "g z = pi" by simp
-
-  show ?thesis
-  proof (cases "z \<in> \<int>")
-    case False
-    with \<open>g z = pi\<close> show ?thesis by (auto simp: g_def divide_simps)
-  next
-    case True
-    then obtain n where n: "z = of_int n" by (elim Ints_cases)
-    with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force
-    moreover have "of_int (1 - n) \<in> \<int>\<^sub>\<le>\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp
-    ultimately show ?thesis using n
-      by (cases "n \<le> 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int)
-  qed
-qed
-
-lemma rGamma_reflection_complex:
-  "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi"
-  using Gamma_reflection_complex[of z]
-    by (simp add: Gamma_def divide_simps split: if_split_asm)
-
-lemma rGamma_reflection_complex':
-  "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi"
-proof -
-  have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))"
-    using rGamma_plus1[of "-z", symmetric] by simp
-  also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi"
-    by (rule rGamma_reflection_complex)
-  finally show ?thesis by simp
-qed
-
-lemma Gamma_reflection_complex':
-  "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))"
-  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac)
-
-
-
-lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi"
-proof -
-  from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1]
-    have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square)
-  hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp
-  also have "\<dots> = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all
-  finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all
-  moreover have "Gamma (1/2 :: real) \<ge> 0" using Gamma_real_pos[of "1/2"] by simp
-  ultimately show ?thesis by (rule real_sqrt_unique [symmetric])
-qed
-
-lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)"
-proof -
-  have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp
-  also have "\<dots> = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real)
-  finally show ?thesis .
-qed
-
-lemma Gamma_legendre_duplication:
-  fixes z :: complex
-  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows "Gamma z * Gamma (z + 1/2) =
-             exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)"
-  using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex)
-
-end
-
-
-subsection \<open>Limits and residues\<close>
-
-text \<open>
-  The inverse of the Gamma function has simple zeros:
-\<close>
-
-lemma rGamma_zeros:
-  "(\<lambda>z. rGamma z / (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n * fact n :: 'a :: Gamma)"
-proof (subst tendsto_cong)
-  let ?f = "\<lambda>z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a"
-  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
-    show "eventually (\<lambda>z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))"
-    by (subst pochhammer_rGamma[of _ "Suc n"])
-       (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0)
-  have "isCont ?f (- of_nat n)" by (intro continuous_intros)
-  thus "?f \<midarrow> (- of_nat n) \<rightarrow> (- 1) ^ n * fact n" unfolding isCont_def
-    by (simp add: pochhammer_same)
-qed
-
-
-text \<open>
-  The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function,
-  and their residues can easily be computed from the limit we have just proven:
-\<close>
-
-lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))"
-proof -
-  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
-    have "eventually (\<lambda>z. rGamma z \<noteq> (0 :: 'a)) (at (- of_nat n))"
-    by (auto elim!: eventually_mono nonpos_Ints_cases'
-             simp: rGamma_eq_zero_iff dist_of_nat dist_minus)
-  with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident]
-    have "filterlim (\<lambda>z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))"
-    unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity])
-                            (simp_all add: filterlim_at)
-  moreover have "(\<lambda>z. inverse (rGamma z) :: 'a) = Gamma"
-    by (intro ext) (simp add: rGamma_inverse_Gamma)
-  ultimately show ?thesis by (simp only: )
-qed
-
-lemma Gamma_residues:
-  "(\<lambda>z. Gamma z * (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n / fact n :: 'a :: Gamma)"
-proof (subst tendsto_cong)
-  let ?c = "(- 1) ^ n / fact n :: 'a"
-  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
-    show "eventually (\<lambda>z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n)))
-            (at (- of_nat n))"
-    by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma)
-  have "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow>
-          inverse ((- 1) ^ n * fact n :: 'a)"
-    by (intro tendsto_intros rGamma_zeros) simp_all
-  also have "inverse ((- 1) ^ n * fact n) = ?c"
-    by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib)
-  finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" .
-qed
-
-
-
-subsection \<open>Alternative definitions\<close>
-
-
-subsubsection \<open>Variant of the Euler form\<close>
-
-
-definition Gamma_series_euler' where
-  "Gamma_series_euler' z n =
-     inverse z * (\<Prod>k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))"
-
-context
-begin
-private lemma Gamma_euler'_aux1:
-  fixes z :: "'a :: {real_normed_field,banach}"
-  assumes n: "n > 0"
-  shows "exp (z * of_real (ln (of_nat n + 1))) = (\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))"
-proof -
-  have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) =
-          exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
-    by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib)
-  also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
-    by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg)
-  also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
-    by (intro setprod.cong) (simp_all add: divide_simps)
-  also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1"
-    by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps)
-  finally show ?thesis ..
-qed
-
-lemma Gamma_series_euler':
-  assumes z: "(z :: 'a :: Gamma) \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows "(\<lambda>n. Gamma_series_euler' z n) \<longlonglongrightarrow> Gamma z"
-proof (rule Gamma_seriesI, rule Lim_transform_eventually)
-  let ?f = "\<lambda>n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)"
-  let ?r = "\<lambda>n. ?f n / Gamma_series z n"
-  let ?r' = "\<lambda>n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))"
-  from z have z': "z \<noteq> 0" by auto
-
-  have "eventually (\<lambda>n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"]
-    using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
-                     elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int)
-  moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))"
-    by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all
-  ultimately show "?r \<longlonglongrightarrow> 1" by (force dest!: Lim_transform_eventually)
-
-  from eventually_gt_at_top[of "0::nat"]
-    show "eventually (\<lambda>n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially"
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    from n z' have "Gamma_series_euler' z n =
-      exp (z * of_real (ln (of_nat n + 1))) / (z * (\<Prod>k=1..n. (1 + z / of_nat k)))"
-      by (subst Gamma_euler'_aux1)
-         (simp_all add: Gamma_series_euler'_def setprod.distrib
-                        setprod_inversef[symmetric] divide_inverse)
-    also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"
-      by (cases n) (simp_all add: pochhammer_setprod fact_setprod atLeastLessThanSuc_atLeastAtMost
-        setprod_dividef [symmetric] field_simps setprod.atLeast_Suc_atMost_Suc_shift)
-    also have "z * \<dots> = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec)
-    finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp
-  qed
-qed
-
-end
-
-
-
-subsubsection \<open>Weierstrass form\<close>
-
-definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
-  "Gamma_series_weierstrass z n =
-     exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))"
-
-definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
-  "rGamma_series_weierstrass z n =
-     exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))"
-
-lemma Gamma_series_weierstrass_nonpos_Ints:
-  "eventually (\<lambda>k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially"
-  using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def)
-
-lemma rGamma_series_weierstrass_nonpos_Ints:
-  "eventually (\<lambda>k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially"
-  using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def)
-
-lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)"
-proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-  case True
-  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
-  also from True have "Gamma_series_weierstrass \<dots> \<longlonglongrightarrow> Gamma z"
-    by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int)
-  finally show ?thesis .
-next
-  case False
-  hence z: "z \<noteq> 0" by auto
-  let ?f = "(\<lambda>x. \<Prod>x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))"
-  have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \<ge> 1" for n :: nat
-    using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp)
-  have "(\<lambda>n. \<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \<longlonglongrightarrow> ln_Gamma z + euler_mascheroni * z + ln z"
-    using ln_Gamma_series'_aux[OF False]
-    by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def
-                   setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
-  from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"
-    by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A)
-  from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z
-    show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z"
-    by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def])
-qed
-
-lemma tendsto_complex_of_real_iff: "((\<lambda>x. complex_of_real (f x)) \<longlongrightarrow> of_real c) F = (f \<longlongrightarrow> c) F"
-  by (rule tendsto_of_real_iff)
-
-lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \<longlonglongrightarrow> Gamma (x :: real)"
-  using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def]
-  by (subst tendsto_complex_of_real_iff [symmetric])
-     (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real)
-
-lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \<longlonglongrightarrow> rGamma (z :: complex)"
-proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-  case True
-  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
-  also from True have "rGamma_series_weierstrass \<dots> \<longlonglongrightarrow> rGamma z"
-    by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int)
-  finally show ?thesis .
-next
-  case False
-  have "rGamma_series_weierstrass z = (\<lambda>n. inverse (Gamma_series_weierstrass z n))"
-    by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def
-                  exp_minus divide_inverse setprod_inversef[symmetric] mult_ac)
-  also from False have "\<dots> \<longlonglongrightarrow> inverse (Gamma z)"
-    by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff)
-  finally show ?thesis by (simp add: Gamma_def)
-qed
-
-subsubsection \<open>Binomial coefficient form\<close>
-
-lemma Gamma_gbinomial:
-  "(\<lambda>n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \<longlonglongrightarrow> rGamma (z+1)"
-proof (cases "z = 0")
-  case False
-  show ?thesis
-  proof (rule Lim_transform_eventually)
-    let ?powr = "\<lambda>a b. exp (b * of_real (ln (of_nat a)))"
-    show "eventually (\<lambda>n. rGamma_series z n / z =
-            ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially"
-    proof (intro always_eventually allI)
-      fix n :: nat
-      from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n"
-        by (simp add: gbinomial_pochhammer' pochhammer_rec)
-      also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z"
-        by (simp add: rGamma_series_def divide_simps exp_minus)
-      finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" ..
-    qed
-
-    from False have "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma z / z" by (intro tendsto_intros)
-    also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z]
-      by (simp add: field_simps)
-    finally show "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma (z+1)" .
-  qed
-qed (simp_all add: binomial_gbinomial [symmetric])
-
-lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)"
-  by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric])
-
-lemma gbinomial_asymptotic:
-  fixes z :: "'a :: Gamma"
-  shows "(\<lambda>n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \<longlonglongrightarrow> 
-           inverse (Gamma (- z))"
-  unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] 
-  by (subst (asm) gbinomial_minus')
-     (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric])
-
-lemma fact_binomial_limit:
-  "(\<lambda>n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \<longlonglongrightarrow> 1 / fact k"
-proof (rule Lim_transform_eventually)
-  have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n))))
-            \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _")
-    using Gamma_gbinomial[of "of_nat k :: 'a"]
-    by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus)
-  also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact)
-  finally show "?f \<longlonglongrightarrow> 1 / fact k" .
-
-  show "eventually (\<lambda>n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0"
-    from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)"
-      by (simp add: exp_of_nat_mult)
-    thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp
-  qed
-qed
-
-lemma binomial_asymptotic':
-  "(\<lambda>n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \<longlonglongrightarrow> 1"
-  using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp
-
-lemma gbinomial_Beta:
-  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))"
-using assms
-proof (induction n arbitrary: z)
-  case 0
-  hence "z + 2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-    using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute)
-  with 0 show ?case
-    by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute)
-next
-  case (Suc n z)
-  show ?case
-  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
-    case True
-    with Suc.prems have "z = 0"
-      by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff)
-    show ?thesis
-    proof (cases "n = 0")
-      case True
-      with \<open>z = 0\<close> show ?thesis
-        by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric])
-    next
-      case False
-      with \<open>z = 0\<close> show ?thesis
-        by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1)
-    qed
-  next
-    case False
-    have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp
-    also have "\<dots> = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)"
-      by (subst gbinomial_factors) (simp add: field_simps)
-    also from False have "\<dots> = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" 
-      (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1)
-    also have "of_nat (Suc n) \<notin> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all
-    hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)"
-      by (subst Beta_plus1_right [symmetric]) simp_all
-    finally show ?thesis .
-  qed
-qed
-
-lemma gbinomial_Gamma:
-  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
-  shows   "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
-proof -
-  have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
-    by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac)
-  also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)"
-    using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac)
-  finally show ?thesis .
-qed
-
-
-subsubsection \<open>Integral form\<close>
-
-lemma integrable_Gamma_integral_bound:
-  fixes a c :: real
-  assumes a: "a > -1" and c: "c \<ge> 0"
-  defines "f \<equiv> \<lambda>x. if x \<in> {0..c} then x powr a else exp (-x/2)"
-  shows   "f integrable_on {0..}"
-proof -
-  have "f integrable_on {0..c}"
-    by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]])
-       (insert a c, simp_all add: f_def)
-  moreover have A: "(\<lambda>x. exp (-x/2)) integrable_on {c..}"
-    using integrable_on_exp_minus_to_infinity[of "1/2"] by simp
-  have "f integrable_on {c..}"
-    by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def)
-  ultimately show "f integrable_on {0..}" 
-    by (rule integrable_union') (insert c, auto simp: max_def)
-qed  
-
-lemma Gamma_integral_complex:
-  assumes z: "Re z > 0"
-  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
-proof -
-  have A: "((\<lambda>t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n))
-          has_integral (fact n / pochhammer z (n+1))) {0..1}"
-    if "Re z > 0" for n z using that
-  proof (induction n arbitrary: z)
-    case 0
-    have "((\<lambda>t. complex_of_real t powr (z - 1)) has_integral
-            (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0
-      by (intro fundamental_theorem_of_calculus_interior)
-         (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex)
-    thus ?case by simp
-  next
-    case (Suc n)
-    let ?f = "\<lambda>t. complex_of_real t powr z / z"
-    let ?f' = "\<lambda>t. complex_of_real t powr (z - 1)"
-    let ?g = "\<lambda>t. (1 - complex_of_real t) ^ Suc n"
-    let ?g' = "\<lambda>t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)"
-    have "((\<lambda>t. ?f' t * ?g t) has_integral
-            (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}"
-      (is "(_ has_integral ?I) _")
-    proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g])
-      from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g"
-        by (auto intro!: continuous_intros)
-    next
-      fix t :: real assume t: "t \<in> {0<..<1}"
-      show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems
-        by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex)
-      show "(?g has_vector_derivative ?g' t) (at t)"
-        by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all
-    next
-      from Suc.prems have [simp]: "z \<noteq> 0" by auto
-      from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp
-      have [simp]: "z + of_nat n \<noteq> 0" "z + 1 + of_nat n \<noteq> 0" for n
-        using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel)
-      have "((\<lambda>x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral
-              fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}"
-        (is "(?A has_integral ?B) _")
-        using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec)
-      also have "?A = (\<lambda>t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps)
-      also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))"
-        by (simp add: divide_simps pochhammer_rec
-              setprod_shift_bounds_cl_Suc_ivl del: of_nat_Suc)
-      finally show "((\<lambda>t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}"
-        by simp
-    qed (simp_all add: bounded_bilinear_mult)
-    thus ?case by simp
-  qed
-
-  have B: "((\<lambda>t. if t \<in> {0..of_nat n} then
-             of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0)
-           has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n
-  proof (cases "n > 0")
-    case [simp]: True
-    hence [simp]: "n \<noteq> 0" by auto
-    with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0]
-      have "((\<lambda>x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n)
-              has_integral fact n * of_nat n / pochhammer z (n+1)) ((\<lambda>x. real n * x)`{0..1})"
-      (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real)
-    also from True have "((\<lambda>x. real n*x)`{0..1}) = {0..real n}"
-      by (subst image_mult_atLeastAtMost) simp_all
-    also have "?f = (\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)"
-      using True by (intro ext) (simp add: field_simps)
-    finally have "((\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
-                    has_integral ?I) {0..real n}" (is ?P) .
-    also have "?P \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n)
-                        has_integral ?I) {0..real n}"
-      by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric])
-    also have "\<dots> \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n)
-                        has_integral ?I) {0..real n}"
-      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div)
-    finally have \<dots> .
-    note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"]
-    have "((\<lambda>x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n)
-            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P)
-      by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric])
-         (simp add: Ln_of_nat algebra_simps)
-    also have "?P \<longleftrightarrow> ((\<lambda>x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
-            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}"
-      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real)
-    also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) =
-                 (of_nat n powr z * fact n / pochhammer z (n+1))"
-      by (auto simp add: powr_def algebra_simps exp_diff)
-    finally show ?thesis by (subst has_integral_restrict) simp_all
-  next
-    case False
-    thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl)
-  qed
-
-  have "eventually (\<lambda>n. Gamma_series z n =
-          of_nat n powr z * fact n / pochhammer z (n+1)) sequentially"
-    using eventually_gt_at_top[of "0::nat"]
-    by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def)
-  from this and Gamma_series_LIMSEQ[of z]
-    have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z"
-    by (rule Lim_transform_eventually)
-
-  {
-    fix x :: real assume x: "x \<ge> 0"
-    have lim_exp: "(\<lambda>k. (1 - x / real k) ^ k) \<longlonglongrightarrow> exp (-x)"
-      using tendsto_exp_limit_sequentially[of "-x"] by simp
-    have "(\<lambda>k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k))
-            \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" (is ?P)
-      by (intro tendsto_intros lim_exp)
-    also from eventually_gt_at_top[of "nat \<lceil>x\<rceil>"]
-      have "eventually (\<lambda>k. of_nat k > x) sequentially" by eventually_elim linarith
-    hence "?P \<longleftrightarrow> (\<lambda>k. if x \<le> of_nat k then
-                 of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0)
-                   \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))"
-      by (intro tendsto_cong) (auto elim!: eventually_mono)
-    finally have \<dots> .
-  }
-  hence D: "\<forall>x\<in>{0..}. (\<lambda>k. if x \<in> {0..real k} then
-              of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0)
-             \<longlonglongrightarrow> of_real x powr (z - 1) / of_real (exp x)"
-    by (simp add: exp_minus field_simps cong: if_cong)
-
-  have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"
-    by (intro tendsto_intros ln_x_over_x_tendsto_0)
-  hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp
-  from order_tendstoD(2)[OF this, of "1/2"]
-    have "eventually (\<lambda>x. (Re z - 1) * ln x / x < 1/2) at_top" by simp
-  from eventually_conj[OF this eventually_gt_at_top[of 0]]
-    obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < 1/2 \<and> x > 0"
-    by (auto simp: eventually_at_top_linorder)
-  hence x0: "x0 > 0" "\<And>x. x \<ge> x0 \<Longrightarrow> (Re z - 1) * ln x < x / 2" by auto
-
-  define h where "h = (\<lambda>x. if x \<in> {0..x0} then x powr (Re z - 1) else exp (-x/2))"
-  have le_h: "x powr (Re z - 1) * exp (-x) \<le> h x" if x: "x \<ge> 0" for x
-  proof (cases "x > x0")
-    case True
-    from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)"
-      by (simp add: powr_def exp_diff exp_minus field_simps exp_add)
-    also from x0(2)[of x] True have "\<dots> < exp (-x/2)"
-      by (simp add: field_simps)
-    finally show ?thesis using True by (auto simp add: h_def)
-  next
-    case False
-    from x have "x powr (Re z - 1) * exp (- x) \<le> x powr (Re z - 1) * 1"
-      by (intro mult_left_mono) simp_all
-    with False show ?thesis by (auto simp add: h_def)
-  qed
-
-  have E: "\<forall>x\<in>{0..}. cmod (if x \<in> {0..real k} then of_real x powr (z - 1) *
-                   (1 - complex_of_real x / of_nat k) ^ k else 0) \<le> h x"
-    (is "\<forall>x\<in>_. ?f x \<le> _") for k
-  proof safe
-    fix x :: real assume x: "x \<ge> 0"
-    {
-      fix x :: real and n :: nat assume x: "x \<le> of_nat n"
-      have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp
-      also have "norm \<dots> = \<bar>(1 - x / real n)\<bar>" by (subst norm_of_real) (rule refl)
-      also from x have "\<dots> = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps)
-      finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" .
-    } note D = this
-    from D[of x k] x
-      have "?f x \<le> (if of_nat k \<ge> x \<and> k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)"
-      by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg)
-    also have "\<dots> \<le> x powr (Re z - 1) * exp  (-x)"
-      by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n)
-    also from x have "\<dots> \<le> h x" by (rule le_h)
-    finally show "?f x \<le> h x" .
-  qed
-  
-  have F: "h integrable_on {0..}" unfolding h_def
-    by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all)
-  show ?thesis
-    by (rule has_integral_dominated_convergence[OF B F E D C])
-qed
-
-lemma Gamma_integral_real:
-  assumes x: "x > (0 :: real)"
-  shows   "((\<lambda>t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}"
-proof -
-  have A: "((\<lambda>t. complex_of_real t powr (complex_of_real x - 1) /
-          complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}"
-    using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real)
-  have "((\<lambda>t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}"
-    by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric])
-  from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def)
-qed
-
-
-
-subsection \<open>The Weierstraß product formula for the sine\<close>
-
-lemma sin_product_formula_complex:
-  fixes z :: complex
-  shows "(\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k^2)) \<longlonglongrightarrow> sin (of_real pi * z)"
-proof -
-  let ?f = "rGamma_series_weierstrass"
-  have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n))
-            \<longlonglongrightarrow> (- of_real pi * inverse z) * (rGamma z * rGamma (- z))"
-    by (intro tendsto_intros rGamma_weierstrass_complex)
-  also have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) =
-                    (\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2))"
-  proof
-    fix n :: nat
-    have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) =
-              of_real pi * z * (\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)"
-      by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus
-                    divide_simps setprod.distrib[symmetric] power2_eq_square)
-    also have "(\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) =
-                 (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)"
-      by (intro setprod.cong) (simp_all add: power2_eq_square field_simps)
-    finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \<dots>"
-      by (simp add: divide_simps)
-  qed
-  also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)"
-    by (subst rGamma_reflection_complex') (simp add: divide_simps)
-  finally show ?thesis .
-qed
-
-lemma sin_product_formula_real:
-  "(\<lambda>n. pi * (x::real) * (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x)"
-proof -
-  from sin_product_formula_complex[of "of_real x"]
-    have "(\<lambda>n. of_real pi * of_real x * (\<Prod>k=1..n. 1 - (of_real x)^2 / (of_nat k)^2))
-              \<longlonglongrightarrow> sin (of_real pi * of_real x :: complex)" (is "?f \<longlonglongrightarrow> ?y") .
-  also have "?f = (\<lambda>n. of_real (pi * x * (\<Prod>k=1..n. 1 - x^2 / (of_nat k^2))))" by simp
-  also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult)
-  finally show ?thesis by (subst (asm) tendsto_of_real_iff)
-qed
-
-lemma sin_product_formula_real':
-  assumes "x \<noteq> (0::real)"
-  shows   "(\<lambda>n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x) / (pi * x)"
-  using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms
-  by simp
-
-
-subsection \<open>The Solution to the Basel problem\<close>
-
-theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)"
-proof -
-  define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n
-  define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)"
-  define f where [abs_def]: "f x = (\<Sum>n. P x n / of_nat (Suc n)^2)" for x
-  define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x
-
-  have sums: "(\<lambda>n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x
-  proof (cases "x = 0")
-    assume x: "x = 0"
-    have "summable (\<lambda>n. inverse ((real_of_nat (Suc n))\<^sup>2))"
-      using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp
-    thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums)
-  next
-    assume x: "x \<noteq> 0"
-    have "(\<lambda>n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))"
-      unfolding P_def using x by (intro telescope_sums' sin_product_formula_real')
-    also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)"
-      unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps)
-    also have "P x 0 = 1" by (simp add: P_def)
-    finally have "(\<lambda>n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" .
-    from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp
-  qed
-
-  have "continuous_on (ball 0 1) f"
-  proof (rule uniform_limit_theorem; (intro always_eventually allI)?)
-    show "uniform_limit (ball 0 1) (\<lambda>n x. \<Sum>k<n. P x k / of_nat (Suc k)^2) f sequentially"
-    proof (unfold f_def, rule weierstrass_m_test)
-      fix n :: nat and x :: real assume x: "x \<in> ball 0 1"
-      {
-        fix k :: nat assume k: "k \<ge> 1"
-        from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1)
-        also from k have "\<dots> \<le> of_nat k^2" by simp
-        finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k
-          by (simp_all add: field_simps del: of_nat_Suc)
-      }
-      hence "(\<Prod>k=1..n. abs (1 - x^2 / of_nat k^2)) \<le> (\<Prod>k=1..n. 1)" by (intro setprod_mono) simp
-      thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2"
-        unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc)
-    qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide)
-  qed (auto simp: P_def intro!: continuous_intros)
-  hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all
-  hence "(f \<midarrow> 0 \<rightarrow> f 0)" by (simp add: isCont_def)
-  also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide)
-  finally have "f \<midarrow> 0 \<rightarrow> K" .
-
-  moreover have "f \<midarrow> 0 \<rightarrow> pi^2 / 6"
-  proof (rule Lim_transform_eventually)
-    define f' where [abs_def]: "f' x = (\<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x
-    have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)"
-      by (auto simp add: eventually_at intro!: exI[of _ 1])
-    thus "eventually (\<lambda>x. f' x = f x) (at 0)"
-    proof eventually_elim
-      fix x :: real assume x: "x \<noteq> 0"
-      have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def)
-      with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"]
-      have "(\<lambda>n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))"
-        by (simp add: eval_nat_numeral)
-      from sums_divide[OF this, of "x^3 * pi"] x
-        have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)"
-        by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac)
-      with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)"
-        by (simp add: g_def)
-      hence "f' x = g x / x^2" by (simp add: sums_iff f'_def)
-      also have "\<dots> = f x" using sums[of x] x by (simp add: sums_iff g_def f_def)
-      finally show "f' x = f x" .
-    qed
-
-    have "isCont f' 0" unfolding f'_def
-    proof (intro isCont_powser_converges_everywhere)
-      fix x :: real show "summable (\<lambda>n. -sin_coeff (n+3) * pi^(n+2) * x^n)"
-      proof (cases "x = 0")
-        assume x: "x \<noteq> 0"
-        from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF
-               sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x
-          show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral)
-      qed (simp only: summable_0_powser)
-    qed
-    hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def)
-    also have "f' 0 = pi * pi / fact 3" unfolding f'_def
-      by (subst powser_zero) (simp add: sin_coeff_def)
-    finally show "f' \<midarrow> 0 \<rightarrow> pi^2 / 6" by (simp add: eval_nat_numeral)
-  qed
-
-  ultimately have "K = pi^2 / 6" by (rule LIM_unique)
-  moreover from inverse_power_summable[of 2]
-    have "summable (\<lambda>n. (inverse (real_of_nat (Suc n)))\<^sup>2)"
-    by (subst summable_Suc_iff) (simp add: power_inverse)
-  ultimately show ?thesis unfolding K_def
-    by (auto simp add: sums_iff power_divide inverse_eq_divide)
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Gamma_Function.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,2969 @@
+(*  Title:    HOL/Multivariate_Analysis/Gamma.thy
+    Author:   Manuel Eberl, TU München
+*)
+
+section \<open>The Gamma Function\<close>
+
+theory Gamma_Function
+imports
+  Complex_Transcendental
+  Summation_Tests
+  Harmonic_Numbers
+  "~~/src/HOL/Library/Nonpos_Ints"
+  "~~/src/HOL/Library/Periodic_Fun"
+begin
+
+text \<open>
+  Several equivalent definitions of the Gamma function and its
+  most important properties. Also contains the definition and some properties
+  of the log-Gamma function and the Digamma function and the other Polygamma functions.
+
+  Based on the Gamma function, we also prove the Weierstraß product form of the
+  sin function and, based on this, the solution of the Basel problem (the
+  sum over all @{term "1 / (n::nat)^2"}.
+\<close>
+
+lemma pochhammer_eq_0_imp_nonpos_Int:
+  "pochhammer (x::'a::field_char_0) n = 0 \<Longrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: pochhammer_eq_0_iff)
+
+lemma closed_nonpos_Ints [simp]: "closed (\<int>\<^sub>\<le>\<^sub>0 :: 'a :: real_normed_algebra_1 set)"
+proof -
+  have "\<int>\<^sub>\<le>\<^sub>0 = (of_int ` {n. n \<le> 0} :: 'a set)"
+    by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int)
+  also have "closed \<dots>" by (rule closed_of_int_image)
+  finally show ?thesis .
+qed
+
+lemma plus_one_in_nonpos_Ints_imp: "z + 1 \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all
+
+lemma of_int_in_nonpos_Ints_iff:
+  "(of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> 0"
+  by (auto simp: nonpos_Ints_def)
+
+lemma one_plus_of_int_in_nonpos_Ints_iff:
+  "(1 + of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> -1"
+proof -
+  have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp
+  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n + 1 \<le> 0" by (subst of_int_in_nonpos_Ints_iff) simp_all
+  also have "\<dots> \<longleftrightarrow> n \<le> -1" by presburger
+  finally show ?thesis .
+qed
+
+lemma one_minus_of_nat_in_nonpos_Ints_iff:
+  "(1 - of_nat n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0"
+proof -
+  have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp
+  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger
+  finally show ?thesis .
+qed
+
+lemma fraction_not_in_ints:
+  assumes "\<not>(n dvd m)" "n \<noteq> 0"
+  shows   "of_int m / of_int n \<notin> (\<int> :: 'a :: {division_ring,ring_char_0} set)"
+proof
+  assume "of_int m / (of_int n :: 'a) \<in> \<int>"
+  then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases)
+  with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps)
+  hence "m = k * n" by (subst (asm) of_int_eq_iff)
+  hence "n dvd m" by simp
+  with assms(1) show False by contradiction
+qed
+
+lemma fraction_not_in_nats:
+  assumes "\<not>n dvd m" "n \<noteq> 0"
+  shows   "of_int m / of_int n \<notin> (\<nat> :: 'a :: {division_ring,ring_char_0} set)"
+proof
+  assume "of_int m / of_int n \<in> (\<nat> :: 'a set)"
+  also note Nats_subset_Ints
+  finally have "of_int m / of_int n \<in> (\<int> :: 'a set)" .
+  moreover have "of_int m / of_int n \<notin> (\<int> :: 'a set)"
+    using assms by (intro fraction_not_in_ints)
+  ultimately show False by contradiction
+qed
+
+lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \<notin> \<int> \<Longrightarrow> z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: Ints_def nonpos_Ints_def)
+
+lemma double_in_nonpos_Ints_imp:
+  assumes "2 * (z :: 'a :: field_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<or> z + 1/2 \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof-
+  from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases')
+  thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps)
+qed
+
+
+lemma sin_series: "(\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
+proof -
+  from sin_converges[of z] have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z" .
+  also have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z \<longleftrightarrow>
+                 (\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
+    by (subst sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
+       (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE)
+  finally show ?thesis .
+qed
+
+lemma cos_series: "(\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
+proof -
+  from cos_converges[of z] have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z" .
+  also have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z \<longleftrightarrow>
+                 (\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
+    by (subst sums_mono_reindex[of "\<lambda>n. 2*n", symmetric])
+       (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE)
+  finally show ?thesis .
+qed
+
+lemma sin_z_over_z_series:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes "z \<noteq> 0"
+  shows   "(\<lambda>n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)"
+proof -
+  from sin_series[of z] have "(\<lambda>n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z"
+    by (simp add: field_simps scaleR_conv_of_real)
+  from sums_mult[OF this, of "inverse z"] and assms show ?thesis
+    by (simp add: field_simps)
+qed
+
+lemma sin_z_over_z_series':
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes "z \<noteq> 0"
+  shows   "(\<lambda>n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)"
+proof -
+  from sums_split_initial_segment[OF sin_converges[of z], of 1]
+    have "(\<lambda>n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp
+  from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps)
+qed
+
+lemma has_field_derivative_sin_z_over_z:
+  fixes A :: "'a :: {real_normed_field,banach} set"
+  shows "((\<lambda>z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)"
+      (is "(?f has_field_derivative ?f') _")
+proof (rule has_field_derivative_at_within)
+  have "((\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n)
+            has_field_derivative (\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)"
+  proof (rule termdiffs_strong)
+    from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1]
+      show "summable (\<lambda>n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def)
+  qed simp
+  also have "(\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f"
+  proof
+    fix z
+    show "(\<Sum>n. of_real (sin_coeff (n+1)) * z^n)  = ?f z"
+      by (cases "z = 0") (insert sin_z_over_z_series'[of z],
+            simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def)
+  qed
+  also have "(\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) =
+                 diffs (\<lambda>n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero)
+  also have "\<dots> = 0" by (simp add: sin_coeff_def diffs_def)
+  finally show "((\<lambda>z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" .
+qed
+
+lemma round_Re_minimises_norm:
+  "norm ((z::complex) - of_int m) \<ge> norm (z - of_int (round (Re z)))"
+proof -
+  let ?n = "round (Re z)"
+  have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)"
+    by (simp add: cmod_def)
+  also have "\<bar>Re z - of_int ?n\<bar> \<le> \<bar>Re z - of_int m\<bar>" by (rule round_diff_minimal)
+  hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \<le> sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)"
+    by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff)
+  also have "\<dots> = norm (z - of_int m)" by (simp add: cmod_def)
+  finally show ?thesis .
+qed
+
+lemma Re_pos_in_ball:
+  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
+  shows   "Re t > 0"
+proof -
+  have "Re (z - t) \<le> norm (z - t)" by (rule complex_Re_le_cmod)
+  also from assms have "\<dots> < Re z / 2" by (simp add: dist_complex_def)
+  finally show "Re t > 0" using assms by simp
+qed
+
+lemma no_nonpos_Int_in_ball_complex:
+  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
+  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases)
+
+lemma no_nonpos_Int_in_ball:
+  assumes "t \<in> ball z (dist z (round (Re z)))"
+  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+proof
+  assume "t \<in> \<int>\<^sub>\<le>\<^sub>0"
+  then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases)
+  have "dist z (of_int n) \<le> dist z t + dist t (of_int n)" by (rule dist_triangle)
+  also from assms have "dist z t < dist z (round (Re z))" by simp
+  also have "\<dots> \<le> dist z (of_int n)"
+    using round_Re_minimises_norm[of z] by (simp add: dist_complex_def)
+  finally have "dist t (of_int n) > 0" by simp
+  with \<open>t = of_int n\<close> show False by simp
+qed
+
+lemma no_nonpos_Int_in_ball':
+  assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  obtains d where "d > 0" "\<And>t. t \<in> ball z d \<Longrightarrow> t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+proof (rule that)
+  from assms show "setdist {z} \<int>\<^sub>\<le>\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto
+next
+  fix t assume "t \<in> ball z (setdist {z} \<int>\<^sub>\<le>\<^sub>0)"
+  thus "t \<notin> \<int>\<^sub>\<le>\<^sub>0" using setdist_le_dist[of z "{z}" t "\<int>\<^sub>\<le>\<^sub>0"] by force
+qed
+
+lemma no_nonpos_Real_in_ball:
+  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0" and t: "t \<in> ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
+  shows   "t \<notin> \<real>\<^sub>\<le>\<^sub>0"
+using z
+proof (cases "Im z = 0")
+  assume A: "Im z = 0"
+  with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff)
+  with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff)
+next
+  assume A: "Im z \<noteq> 0"
+  have "abs (Im z) - abs (Im t) \<le> abs (Im z - Im t)" by linarith
+  also have "\<dots> = abs (Im (z - t))" by simp
+  also have "\<dots> \<le> norm (z - t)" by (rule abs_Im_le_cmod)
+  also from A t have "\<dots> \<le> abs (Im z) / 2" by (simp add: dist_complex_def)
+  finally have "abs (Im t) > 0" using A by simp
+  thus ?thesis by (force simp add: complex_nonpos_Reals_iff)
+qed
+
+
+subsection \<open>Definitions\<close>
+
+text \<open>
+  We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}.
+  This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its
+  properties more convenient because one does not have to watch out for discontinuities.
+  (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere,
+  whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers)
+
+  We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage
+  that it is a relatively simple limit that converges everywhere. The limit at the poles is 0
+  (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows
+  immediately from the definition.
+\<close>
+
+definition Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)"
+
+definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n"
+
+definition rGamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))"
+
+lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)"
+  and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)"
+  unfolding Gamma_series_def rGamma_series_def by simp_all
+
+lemma rGamma_series_minus_of_nat:
+  "eventually (\<lambda>n. rGamma_series (- of_nat k) n = 0) sequentially"
+  using eventually_ge_at_top[of k]
+  by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff)
+
+lemma Gamma_series_minus_of_nat:
+  "eventually (\<lambda>n. Gamma_series (- of_nat k) n = 0) sequentially"
+  using eventually_ge_at_top[of k]
+  by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff)
+
+lemma Gamma_series'_minus_of_nat:
+  "eventually (\<lambda>n. Gamma_series' (- of_nat k) n = 0) sequentially"
+  using eventually_gt_at_top[of k]
+  by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff)
+
+lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma_series z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp)
+
+lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp)
+
+lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series' z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp)
+
+lemma Gamma_series_Gamma_series':
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(\<lambda>n. Gamma_series' z n / Gamma_series z n) \<longlonglongrightarrow> 1"
+proof (rule Lim_transform_eventually)
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n"
+      by (cases n, simp)
+         (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec'
+               dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp)
+    also from n have "\<dots> = z / of_nat n + 1" by (simp add: divide_simps)
+    finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" ..
+  qed
+  have "(\<lambda>x. z / of_nat x) \<longlonglongrightarrow> 0"
+    by (rule tendsto_norm_zero_cancel)
+       (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n],
+        simp add:  norm_divide inverse_eq_divide)
+  from tendsto_add[OF this tendsto_const[of 1]] show "(\<lambda>n. z / of_nat n + 1) \<longlonglongrightarrow> 1" by simp
+qed
+
+
+subsection \<open>Convergence of the Euler series form\<close>
+
+text \<open>
+  We now show that the series that defines the Gamma function in the Euler form converges
+  and that the function defined by it is continuous on the complex halfspace with positive
+  real part.
+
+  We do this by showing that the logarithm of the Euler series is continuous and converges
+  locally uniformly, which means that the log-Gamma function defined by its limit is also
+  continuous.
+
+  This will later allow us to lift holomorphicity and continuity from the log-Gamma
+  function to the inverse of the Gamma function, and from that to the Gamma function itself.
+\<close>
+
+definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))"
+
+definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "ln_Gamma_series' z n =
+     - euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))"
+
+definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> 'a" where
+  "ln_Gamma z = lim (ln_Gamma_series z)"
+
+
+text \<open>
+  We now show that the log-Gamma series converges locally uniformly for all complex numbers except
+  the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this
+  proof:
+  http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm
+\<close>
+
+context
+begin
+
+private lemma ln_Gamma_series_complex_converges_aux:
+  fixes z :: complex and k :: nat
+  assumes z: "z \<noteq> 0" and k: "of_nat k \<ge> 2*norm z" "k \<ge> 2"
+  shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \<le> 2*(norm z + norm z^2) / of_nat k^2"
+proof -
+  let ?k = "of_nat k :: complex" and ?z = "norm z"
+  have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)"
+    by (simp add: algebra_simps)
+  also have "norm ... \<le> ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)"
+    by (subst norm_mult [symmetric], rule norm_triangle_ineq)
+  also have "norm (Ln (1 + -1/?k) - (-1/?k)) \<le> (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))"
+    using k by (intro Ln_approx_linear) (simp add: norm_divide)
+  hence "?z * norm (ln (1-1/?k) + 1/?k) \<le> ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))"
+    by (intro mult_left_mono) simp_all
+  also have "... \<le> (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k
+    by (simp add: field_simps power2_eq_square norm_divide)
+  also have "... \<le> (?z * 2) / of_nat k^2" using k
+    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
+  also have "norm (ln (1+z/?k) - z/?k) \<le> norm (z/?k)^2 / (1 - norm (z/?k))" using k
+    by (intro Ln_approx_linear) (simp add: norm_divide)
+  hence "norm (ln (1+z/?k) - z/?k) \<le> ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)"
+    by (simp add: field_simps norm_divide)
+  also have "... \<le> (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k
+    by (simp add: field_simps power2_eq_square)
+  also have "... \<le> (?z^2 * 2) / of_nat k^2" using k
+    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
+  also note add_divide_distrib [symmetric]
+  finally show ?thesis by (simp only: distrib_left mult.commute)
+qed
+
+lemma ln_Gamma_series_complex_converges:
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  assumes d: "d > 0" "\<And>n. n \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> norm (z - of_int n) > d"
+  shows "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n :: complex)"
+proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI')
+  fix e :: real assume e: "e > 0"
+  define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)"
+  define e' where "e' = e / (2*e'')"
+  have "bounded ((\<lambda>t. norm t + norm t^2) ` cball z d)"
+    by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros)
+  hence "bounded ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto
+  hence bdd: "bdd_above ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above)
+
+  with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def
+    by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z])
+  have e'': "norm t + norm t^2 \<le> e''" if "t \<in> ball z d" for t unfolding e''_def using that
+    by (rule cSUP_upper[OF _ bdd])
+  from e z e''_pos have e': "e' > 0" unfolding e'_def
+    by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps)
+
+  have "summable (\<lambda>k. inverse ((real_of_nat k)^2))"
+    by (rule inverse_power_summable) simp
+  from summable_partial_sum_bound[OF this e'] guess M . note M = this
+
+  define N where "N = max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
+  {
+    from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>"
+      by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
+    hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def
+      by (simp_all add: le_of_int_ceiling)
+    also have "... \<le> of_nat N" unfolding N_def
+      by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1)
+    finally have "of_nat N \<ge> 2 * (norm z + d)" .
+    moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all
+    moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n
+      using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def
+      by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps)
+    moreover note calculation
+  } note N = this
+
+  show "\<exists>M. \<forall>t\<in>ball z d. \<forall>m\<ge>M. \<forall>n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e"
+    unfolding dist_complex_def
+  proof (intro exI[of _ N] ballI allI impI)
+    fix t m n assume t: "t \<in> ball z d" and mn: "m \<ge> N" "n > m"
+    from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def)
+    also have "dist z 0 - dist z t \<le> dist 0 t" using dist_triangle[of 0 z t]
+      by (simp add: dist_commute)
+    finally have t_nz: "t \<noteq> 0" by auto
+
+    have "norm t \<le> norm z + norm (t - z)" by (rule norm_triangle_sub)
+    also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute)
+    also have "2 * (norm z + d) \<le> of_nat N" by (rule N)
+    also have "N \<le> m" by (rule mn)
+    finally have norm_t: "2 * norm t < of_nat m" by simp
+
+    have "ln_Gamma_series t m - ln_Gamma_series t n =
+              (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) +
+              ((\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)))"
+      by (simp add: ln_Gamma_series_def algebra_simps)
+    also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) =
+                 (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
+      by (simp add: setsum_diff)
+    also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce
+    also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) =
+                   (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn
+      by (subst setsum_telescope'' [symmetric]) simp_all
+    also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N
+      by (intro setsum_cong_Suc)
+         (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat)
+    also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k
+      using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps)
+    hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) =
+             (\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N
+      by (intro setsum.cong) simp_all
+    also note setsum.distrib [symmetric]
+    also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le>
+      (\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t
+      by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
+    also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
+      by (simp add: setsum_right_distrib)
+    also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
+      by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
+    also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')"
+      by (simp add: e'_def field_simps power2_eq_square)
+    also from e''[OF t] e''_pos e
+      have "\<dots> \<le> e * 1" by (intro mult_left_mono) (simp_all add: field_simps)
+    finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp
+  qed
+qed
+
+end
+
+lemma ln_Gamma_series_complex_converges':
+  assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
+proof -
+  define d' where "d' = Re z"
+  define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)"
+  have "of_int (round d') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 0" using that
+    by (intro nonpos_Ints_of_int) (simp_all add: round_def)
+  with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less)
+
+  have "d < cmod (z - of_int n)" if "n \<in> \<int>\<^sub>\<le>\<^sub>0" for n
+  proof (cases "Re z > 0")
+    case True
+    from nonpos_Ints_nonpos[OF that] have n: "n \<le> 0" by simp
+    from True have "d = Re z/2" by (simp add: d_def d'_def)
+    also from n True have "\<dots> < Re (z - of_int n)" by simp
+    also have "\<dots> \<le> norm (z - of_int n)" by (rule complex_Re_le_cmod)
+    finally show ?thesis .
+  next
+    case False
+    with assms nonpos_Ints_of_int[of "round (Re z)"]
+      have "z \<noteq> of_int (round d')" by (auto simp: not_less)
+    with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def)
+    also have "\<dots> \<le> norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm)
+    finally show ?thesis .
+  qed
+  hence conv: "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
+    by (intro ln_Gamma_series_complex_converges d_pos z) simp_all
+  from d_pos conv show ?thesis by blast
+qed
+
+lemma ln_Gamma_series_complex_converges'': "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> convergent (ln_Gamma_series z)"
+  by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent)
+
+lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z"
+  using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def)
+
+lemma exp_ln_Gamma_series_complex:
+  assumes "n > 0" "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "exp (ln_Gamma_series z n :: complex) = Gamma_series z n"
+proof -
+  from assms obtain m where m: "n = Suc m" by (cases n) blast
+  from assms have "z \<noteq> 0" by (intro notI) auto
+  with assms have "exp (ln_Gamma_series z n) =
+          (of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"
+    unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum)
+  also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)"
+    by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp)
+  also have "... = (\<Prod>k=1..n. z + k) / fact n"
+    by (simp add: fact_setprod)
+    (subst setprod_dividef [symmetric], simp_all add: field_simps)
+  also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n"
+    by (simp add: setprod.atLeast0_atMost_Suc_shift setprod.atLeast_Suc_atMost_Suc_shift)
+  also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)"
+    unfolding pochhammer_setprod
+    by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
+  also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n"
+    unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat)
+  finally show ?thesis .
+qed
+
+
+lemma ln_Gamma_series'_aux:
+  assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(\<lambda>k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums
+              (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s")
+unfolding sums_def
+proof (rule Lim_transform)
+  show "(\<lambda>n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \<longlonglongrightarrow> ?s"
+    (is "?g \<longlonglongrightarrow> _")
+    by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms)
+
+  have A: "eventually (\<lambda>n. (\<Sum>k<n. ?f k) - ?g n = 0) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"
+      by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric],
+          subst atLeastLessThanSuc_atLeastAtMost) simp_all
+    also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
+      by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse)
+    also from n have "\<dots> - ?g n = 0"
+      by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat)
+    finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
+  qed
+  show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
+qed
+
+
+lemma uniformly_summable_deriv_ln_Gamma:
+  assumes z: "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" and d: "d > 0" "d \<le> norm z/2"
+  shows "uniformly_convergent_on (ball z d)
+            (\<lambda>k z. \<Sum>i<k. inverse (of_nat (Suc i)) - inverse (z + of_nat (Suc i)))"
+           (is "uniformly_convergent_on _ (\<lambda>k z. \<Sum>i<k. ?f i z)")
+proof (rule weierstrass_m_test'_ev)
+  {
+    fix t assume t: "t \<in> ball z d"
+    have "norm z = norm (t + (z - t))" by simp
+    have "norm (t + (z - t)) \<le> norm t + norm (z - t)" by (rule norm_triangle_ineq)
+    also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm)
+    finally have A: "norm t > norm z / 2" using z by (simp add: field_simps)
+
+    have "norm t = norm (z + (t - z))" by simp
+    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
+    also from t d have "norm (t - z) \<le> norm z / 2" by (simp add: dist_norm norm_minus_commute)
+    also from z have "\<dots> < norm z" by simp
+    finally have B: "norm t < 2 * norm z" by simp
+    note A B
+  } note ball = this
+
+  show "eventually (\<lambda>n. \<forall>t\<in>ball z d. norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially"
+    using eventually_gt_at_top apply eventually_elim
+  proof safe
+    fix t :: 'a assume t: "t \<in> ball z d"
+    from z ball[OF t] have t_nz: "t \<noteq> 0" by auto
+    fix n :: nat assume n: "n > nat \<lceil>4 * norm z\<rceil>"
+    from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp
+    also from n have "\<dots>  < of_nat n" by linarith
+    finally have n: "of_nat n > 2 * norm t" .
+    hence "of_nat n > norm t" by simp
+    hence t': "t \<noteq> -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc)
+
+    with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))"
+      by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc)
+    also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))"
+      by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc)
+    also {
+      from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)"
+        by (intro divide_left_mono mult_pos_pos) simp_all
+      also have "\<dots> < norm (of_nat (Suc n) / t) - norm (1 :: 'a)"
+        using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc)
+      also have "\<dots> \<le> norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq)
+      finally have "inverse (norm (of_nat (Suc n)/t + 1)) \<le> 4 * norm z / of_nat (Suc n)"
+        using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc)
+    }
+    also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) =
+                 4 * norm z * inverse (of_nat (Suc n)^2)"
+                 by (simp add: divide_simps power2_eq_square del: of_nat_Suc)
+    finally show "norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)"
+      by (simp del: of_nat_Suc)
+  qed
+next
+  show "summable (\<lambda>n. 4 * norm z * inverse ((of_nat (Suc n))^2))"
+    by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable)
+qed
+
+lemma summable_deriv_ln_Gamma:
+  "z \<noteq> (0 :: 'a :: {real_normed_field,banach}) \<Longrightarrow>
+     summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))"
+  unfolding summable_iff_convergent
+  by (rule uniformly_convergent_imp_convergent,
+      rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all
+
+
+definition Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+  "Polygamma n z = (if n = 0 then
+      (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else
+      (-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))"
+
+abbreviation Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+  "Digamma \<equiv> Polygamma 0"
+
+lemma Digamma_def:
+  "Digamma z = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni"
+  by (simp add: Polygamma_def)
+
+
+lemma summable_Digamma:
+  assumes "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0"
+  shows   "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
+proof -
+  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
+                       (0 - inverse (z + of_nat 0))"
+    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+  from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]]
+    show "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp
+qed
+
+lemma summable_offset:
+  assumes "summable (\<lambda>n. f (n + k) :: 'a :: real_normed_vector)"
+  shows   "summable f"
+proof -
+  from assms have "convergent (\<lambda>m. \<Sum>n<m. f (n + k))" by (simp add: summable_iff_convergent)
+  hence "convergent (\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)))"
+    by (intro convergent_add convergent_const)
+  also have "(\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k))) = (\<lambda>m. \<Sum>n<m+k. f n)"
+  proof
+    fix m :: nat
+    have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto
+    also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)"
+      by (rule setsum.union_disjoint) auto
+    also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
+      by (intro setsum.reindex_cong[of "\<lambda>n. n + k"])
+         (simp, subst image_add_atLeastLessThan, auto)
+    finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan)
+  qed
+  finally have "(\<lambda>a. setsum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. setsum f {..<m + k})"
+    by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset)
+  thus ?thesis by (auto simp: summable_iff_convergent convergent_def)
+qed
+
+lemma Polygamma_converges:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
+  shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))"
+proof (rule weierstrass_m_test'_ev)
+  define e where "e = (1 + d / norm z)"
+  define m where "m = nat \<lceil>norm z * e\<rceil>"
+  {
+    fix t assume t: "t \<in> ball z d"
+    have "norm t = norm (z + (t - z))" by simp
+    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
+    also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute)
+    finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def)
+  } note ball = this
+
+  show "eventually (\<lambda>k. \<forall>t\<in>ball z d. norm (inverse ((t + of_nat k)^n)) \<le>
+            inverse (of_nat (k - m)^n)) sequentially"
+    using eventually_gt_at_top[of m] apply eventually_elim
+  proof (intro ballI)
+    fix k :: nat and t :: 'a assume k: "k > m" and t: "t \<in> ball z d"
+    from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff)
+    also have "\<dots> \<le> norm (of_nat k :: 'a) - norm z * e"
+      unfolding m_def by (subst norm_of_nat) linarith
+    also from ball[OF t] have "\<dots> \<le> norm (of_nat k :: 'a) - norm t" by simp
+    also have "\<dots> \<le> norm (of_nat k + t)" by (rule norm_diff_ineq)
+    finally have "inverse ((norm (t + of_nat k))^n) \<le> inverse (real_of_nat (k - m)^n)" using k n
+      by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc)
+    thus "norm (inverse ((t + of_nat k)^n)) \<le> inverse (of_nat (k - m)^n)"
+      by (simp add: norm_inverse norm_power power_inverse)
+  qed
+
+  have "summable (\<lambda>k. inverse ((real_of_nat k)^n))"
+    using inverse_power_summable[of n] n by simp
+  hence "summable (\<lambda>k. inverse ((real_of_nat (k + m - m))^n))" by simp
+  thus "summable (\<lambda>k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset)
+qed
+
+lemma Polygamma_converges':
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
+  shows "summable (\<lambda>k. inverse ((z + of_nat k)^n))"
+  using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z]
+  by (simp add: summable_iff_convergent)
+
+lemma has_field_derivative_ln_Gamma_complex [derivative_intros]:
+  fixes z :: complex
+  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
+  shows   "(ln_Gamma has_field_derivative Digamma z) (at z)"
+proof -
+  have not_nonpos_Int [simp]: "t \<notin> \<int>\<^sub>\<le>\<^sub>0" if "Re t > 0" for t
+    using that by (auto elim!: nonpos_Ints_cases')
+  from z have z': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" and z'': "z \<noteq> 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I
+     by blast+
+  let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))"
+  let ?f = "\<lambda>z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\<lambda>z. \<Sum>n. ?f' z n"
+  define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
+  from z have d: "d > 0" "norm z/2 \<ge> d" by (auto simp add: complex_nonpos_Reals_iff d_def)
+  have ball: "Im t = 0 \<longrightarrow> Re t > 0" if "dist z t < d" for t
+    using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff)
+  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
+                       (0 - inverse (z + of_nat 0))"
+    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+
+  have "((\<lambda>z. \<Sum>n. ?f z n) has_field_derivative ?F' z) (at z)"
+    using d z ln_Gamma_series'_aux[OF z']
+    apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma)
+    apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball
+             simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff
+             simp del: of_nat_Suc)
+    apply (auto simp add: complex_nonpos_Reals_iff)
+    done
+  with z have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative
+                   ?F' z - euler_mascheroni - inverse z) (at z)"
+    by (force intro!: derivative_eq_intros simp: Digamma_def)
+  also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp
+  also from sums have "-inverse z = (\<Sum>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))"
+    by (simp add: sums_iff)
+  also from sums summable_deriv_ln_Gamma[OF z'']
+    have "?F' z + \<dots> =  (\<Sum>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
+    by (subst suminf_add) (simp_all add: add_ac sums_iff)
+  also have "\<dots> - euler_mascheroni = Digamma z" by (simp add: Digamma_def)
+  finally have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z)
+                    has_field_derivative Digamma z) (at z)" .
+  moreover from eventually_nhds_ball[OF d(1), of z]
+    have "eventually (\<lambda>z. ln_Gamma z = (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)"
+  proof eventually_elim
+    fix t assume "t \<in> ball z d"
+    hence "t \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases)
+    from ln_Gamma_series'_aux[OF this]
+      show "ln_Gamma t = (\<Sum>k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff)
+  qed
+  ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
+qed
+
+declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros]
+
+
+lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni"
+  by (simp add: Digamma_def)
+
+lemma Digamma_plus1:
+  assumes "z \<noteq> 0"
+  shows   "Digamma (z+1) = Digamma z + 1/z"
+proof -
+  have sums: "(\<lambda>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))
+                  sums (inverse (z + of_nat 0) - 0)"
+    by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+  have "Digamma (z+1) = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) -
+          euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac)
+  also have "suminf ?f = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) +
+                         (\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))"
+    using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff)
+  also have "(\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z"
+    using sums by (simp add: sums_iff inverse_eq_divide)
+  finally show ?thesis by (simp add: Digamma_def[of z])
+qed
+
+lemma Polygamma_plus1:
+  assumes "z \<noteq> 0"
+  shows   "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)"
+proof (cases "n = 0")
+  assume n: "n \<noteq> 0"
+  let ?f = "\<lambda>k. inverse ((z + of_nat k) ^ Suc n)"
+  have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\<Sum>k. ?f (k+1))"
+    using n by (simp add: Polygamma_def add_ac)
+  also have "(\<Sum>k. ?f (k+1)) + (\<Sum>k<1. ?f k) = (\<Sum>k. ?f k)"
+    using Polygamma_converges'[OF assms, of "Suc n"] n
+    by (subst suminf_split_initial_segment [symmetric]) simp_all
+  hence "(\<Sum>k. ?f (k+1)) = (\<Sum>k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps)
+  also have "(-1) ^ Suc n * fact n * ((\<Sum>k. ?f k) - inverse (z ^ Suc n)) =
+               Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n
+    by (simp add: inverse_eq_divide algebra_simps Polygamma_def)
+  finally show ?thesis .
+qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide)
+
+lemma Digamma_of_nat:
+  "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni"
+proof (induction n)
+  case (Suc n)
+  have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp
+  also have "\<dots> = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))"
+    by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc)
+  also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc)
+  also have "\<dots> + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni"
+    by (simp add: harm_Suc)
+  finally show ?case .
+qed (simp add: harm_def)
+
+lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni"
+  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl)
+
+lemma Polygamma_of_real: "x \<noteq> 0 \<Longrightarrow> Polygamma n (of_real x) = of_real (Polygamma n x)"
+  unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"]
+  by (simp_all add: suminf_of_real)
+
+lemma Polygamma_Real: "z \<in> \<real> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Polygamma n z \<in> \<real>"
+  by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all
+
+lemma Digamma_half_integer:
+  "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) =
+      (\<Sum>k<n. 2 / (of_nat (2*k+1))) - euler_mascheroni - of_real (2 * ln 2)"
+proof (induction n)
+  case 0
+  have "Digamma (1/2 :: 'a) = of_real (Digamma (1/2))" by (simp add: Polygamma_of_real [symmetric])
+  also have "Digamma (1/2::real) =
+               (\<Sum>k. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni"
+    by (simp add: Digamma_def add_ac)
+  also have "(\<Sum>k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) =
+             (\<Sum>k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))"
+    by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide)
+  also have "\<dots> = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums']
+    by (subst suminf_mult) (simp_all add: algebra_simps sums_iff)
+  finally show ?case by simp
+next
+  case (Suc n)
+  have nz: "2 * of_nat n + (1:: 'a) \<noteq> 0"
+     using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac)
+  hence nz': "of_nat n + (1/2::'a) \<noteq> 0" by (simp add: field_simps)
+  have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp
+  also from nz' have "\<dots> = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)"
+    by (rule Digamma_plus1)
+  also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)"
+    by (subst divide_eq_eq) simp_all
+  also note Suc
+  finally show ?case by (simp add: add_ac)
+qed
+
+lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)"
+  using Digamma_half_integer[of 0] by simp
+
+lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0"
+proof -
+  have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp
+  also have "\<dots> = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp
+  also note euler_mascheroni_less_13_over_22
+  also note ln2_le_25_over_36
+  finally show ?thesis by simp
+qed
+
+
+lemma has_field_derivative_Polygamma [derivative_intros]:
+  fixes z :: "'a :: {real_normed_field,euclidean_space}"
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)"
+proof (rule has_field_derivative_at_within, cases "n = 0")
+  assume n: "n = 0"
+  let ?f = "\<lambda>k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)"
+  let ?F = "\<lambda>z. \<Sum>k. ?f k z" and ?f' = "\<lambda>k z. inverse ((z + of_nat k)\<^sup>2)"
+  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
+  from z have summable: "summable (\<lambda>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))"
+    by (intro summable_Digamma) force
+  from z have conv: "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)\<^sup>2))"
+    by (intro Polygamma_converges) auto
+  with d have "summable (\<lambda>k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent
+    by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent )
+
+  have "(?F has_field_derivative (\<Sum>k. ?f' k z)) (at z)"
+  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
+    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
+    from t d(2)[of t] show "((\<lambda>z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)"
+      by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc
+               dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases)
+  qed (insert d(1) summable conv, (assumption|simp)+)
+  with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
+    unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n
+    by (force simp: power2_eq_square intro!: derivative_eq_intros)
+next
+  assume n: "n \<noteq> 0"
+  from z have z': "z \<noteq> 0" by auto
+  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
+  define n' where "n' = Suc n"
+  from n have n': "n' \<ge> 2" by (simp add: n'_def)
+  have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
+                (\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)"
+  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
+    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
+    with d have t': "t \<notin> \<int>\<^sub>\<le>\<^sub>0" "t \<noteq> 0" by auto
+    show "((\<lambda>a. inverse ((a + of_nat k) ^ n')) has_field_derivative
+                - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t'
+      by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp)
+  next
+    have "uniformly_convergent_on (ball z d)
+              (\<lambda>k z. (- of_nat n' :: 'a) * (\<Sum>i<k. inverse ((z + of_nat i) ^ (n'+1))))"
+      using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
+    thus "uniformly_convergent_on (ball z d)
+              (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
+      by (subst (asm) setsum_right_distrib) simp
+  qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
+  also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
+               (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
+    using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all
+  finally have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
+                    - of_nat n' * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" .
+  from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"]
+    show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
+    unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps)
+qed
+
+declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros]
+
+lemma isCont_Polygamma [continuous_intros]:
+  fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
+  shows "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Polygamma n (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Polygamma]])
+
+lemma continuous_on_Polygamma:
+  "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A (Polygamma n :: _ \<Rightarrow> 'a :: {real_normed_field,euclidean_space})"
+  by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast
+
+lemma isCont_ln_Gamma_complex [continuous_intros]:
+  fixes f :: "'a::t2_space \<Rightarrow> complex"
+  shows "isCont f z \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>z. ln_Gamma (f z)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]])
+
+lemma continuous_on_ln_Gamma_complex [continuous_intros]:
+  fixes A :: "complex set"
+  shows "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A ln_Gamma"
+  by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident])
+     fastforce
+
+text \<open>
+  We define a type class that captures all the fundamental properties of the inverse of the Gamma function
+  and defines the Gamma function upon that. This allows us to instantiate the type class both for
+  the reals and for the complex numbers with a minimal amount of proof duplication.
+\<close>
+
+class Gamma = real_normed_field + complete_space +
+  fixes rGamma :: "'a \<Rightarrow> 'a"
+  assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
+  assumes differentiable_rGamma_aux1:
+    "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
+     let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
+               \<longlonglongrightarrow> d) - scaleR euler_mascheroni 1
+     in  filterlim (\<lambda>y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R
+                        norm (y - z)) (nhds 0) (at z)"
+  assumes differentiable_rGamma_aux2:
+    "let z = - of_nat n
+     in  filterlim (\<lambda>y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R
+                        norm (y - z)) (nhds 0) (at z)"
+  assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
+             let fact' = (\<lambda>n. setprod of_nat {1..n});
+                 exp = (\<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x^k /\<^sub>R fact k) \<longlonglongrightarrow> e);
+                 pochhammer' = (\<lambda>a n. (\<Prod>n = 0..n. a + of_nat n))
+             in  filterlim (\<lambda>n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1))))
+                     (nhds (rGamma z)) sequentially"
+begin
+subclass banach ..
+end
+
+definition "Gamma z = inverse (rGamma z)"
+
+
+subsection \<open>Basic properties\<close>
+
+lemma Gamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z = 0"
+  and rGamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z = 0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma Gamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z \<noteq> 0"
+  and rGamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z \<noteq> 0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma Gamma_eq_zero_iff: "Gamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+  and rGamma_eq_zero_iff: "rGamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)"
+  unfolding Gamma_def by simp
+
+lemma rGamma_series_LIMSEQ [tendsto_intros]:
+  "rGamma_series z \<longlonglongrightarrow> rGamma z"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case False
+  hence "z \<noteq> - of_nat n" for n by auto
+  from rGamma_series_aux[OF this] show ?thesis
+    by (simp add: rGamma_series_def[abs_def] fact_setprod pochhammer_Suc_setprod
+                  exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost)
+qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ)
+
+lemma Gamma_series_LIMSEQ [tendsto_intros]:
+  "Gamma_series z \<longlonglongrightarrow> Gamma z"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case False
+  hence "(\<lambda>n. inverse (rGamma_series z n)) \<longlonglongrightarrow> inverse (rGamma z)"
+    by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff)
+  also have "(\<lambda>n. inverse (rGamma_series z n)) = Gamma_series z"
+    by (simp add: rGamma_series_def Gamma_series_def[abs_def])
+  finally show ?thesis by (simp add: Gamma_def)
+qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ)
+
+lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)"
+  using Gamma_series_LIMSEQ[of z] by (simp add: limI)
+
+lemma rGamma_1 [simp]: "rGamma 1 = 1"
+proof -
+  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
+                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
+  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
+  moreover have "rGamma_series 1 \<longlonglongrightarrow> rGamma 1" by (rule tendsto_intros)
+  ultimately show ?thesis by (intro LIMSEQ_unique)
+qed
+
+lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z"
+proof -
+  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
+  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
+             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
+      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
+    also from n have "\<dots> = ?f n * rGamma_series z n"
+      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
+    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
+  qed
+  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
+    by (intro tendsto_intros lim_inverse_n)
+  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
+  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
+    by (rule Lim_transform_eventually)
+  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
+    by (intro tendsto_intros)
+  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
+qed
+
+
+lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)"
+proof (induction n arbitrary: z)
+  case (Suc n z)
+  have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH)
+  also note rGamma_plus1 [symmetric]
+  finally show ?case by (simp add: add_ac pochhammer_rec')
+qed simp_all
+
+lemma Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z"
+  using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff)
+
+lemma pochhammer_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> pochhammer z n = Gamma (z + of_nat n) / Gamma z"
+  using pochhammer_rGamma[of z]
+  by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps)
+
+lemma Gamma_0 [simp]: "Gamma 0 = 0"
+  and rGamma_0 [simp]: "rGamma 0 = 0"
+  and Gamma_neg_1 [simp]: "Gamma (- 1) = 0"
+  and rGamma_neg_1 [simp]: "rGamma (- 1) = 0"
+  and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0"
+  and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0"
+  and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0"
+  and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0"
+  by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff)
+
+lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp
+
+lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n"
+  by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff
+        of_nat_Suc [symmetric] del: of_nat_Suc)
+
+lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)"
+  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc,
+      subst of_nat_Suc, subst Gamma_fact) (rule refl)
+
+lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)"
+proof (cases "n > 0")
+  case True
+  hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all
+  with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp
+qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int)
+
+lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)"
+  by (simp add: Gamma_of_int rGamma_inverse_Gamma)
+
+lemma Gamma_seriesI:
+  assumes "(\<lambda>n. g n / Gamma_series z n) \<longlonglongrightarrow> 1"
+  shows   "g \<longlonglongrightarrow> Gamma z"
+proof (rule Lim_transform_eventually)
+  have "1/2 > (0::real)" by simp
+  from tendstoD[OF assms, OF this]
+    show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially"
+    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+  from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z"
+    by (intro tendsto_intros)
+  thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp
+qed
+
+lemma Gamma_seriesI':
+  assumes "f \<longlonglongrightarrow> rGamma z"
+  assumes "(\<lambda>n. g n * f n) \<longlonglongrightarrow> 1"
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "g \<longlonglongrightarrow> Gamma z"
+proof (rule Lim_transform_eventually)
+  have "1/2 > (0::real)" by simp
+  from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially"
+    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+  from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z"
+    by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff)
+  thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse)
+qed
+
+lemma Gamma_series'_LIMSEQ: "Gamma_series' z \<longlonglongrightarrow> Gamma z"
+  by (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series']
+                                      Gamma_series'_nonpos_Ints_LIMSEQ[of z])
+
+
+subsection \<open>Differentiability\<close>
+
+lemma has_field_derivative_rGamma_no_nonpos_int:
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)"
+proof (rule has_field_derivative_at_within)
+  from assms have "z \<noteq> - of_nat n" for n by auto
+  from differentiable_rGamma_aux1[OF this]
+    show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)"
+         unfolding Digamma_def suminf_def sums_def[abs_def]
+                   has_field_derivative_def has_derivative_def netlimit_at
+    by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric])
+qed
+
+lemma has_field_derivative_rGamma_nonpos_int:
+  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)"
+  apply (rule has_field_derivative_at_within)
+  using differentiable_rGamma_aux2[of n]
+  unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at
+  by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_setprod) simp
+
+lemma has_field_derivative_rGamma [derivative_intros]:
+  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>norm z\<rfloor>) * fact (nat \<lfloor>norm z\<rfloor>)
+      else -rGamma z * Digamma z)) (at z within A)"
+using has_field_derivative_rGamma_no_nonpos_int[of z A]
+      has_field_derivative_rGamma_nonpos_int[of "nat \<lfloor>norm z\<rfloor>" A]
+  by (auto elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros]
+declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros]
+declare has_field_derivative_rGamma_nonpos_int [derivative_intros]
+declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros]
+declare has_field_derivative_rGamma [derivative_intros]
+
+
+lemma has_field_derivative_Gamma [derivative_intros]:
+  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)"
+  unfolding Gamma_def [abs_def]
+  by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff)
+
+declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros]
+
+(* TODO: Hide ugly facts properly *)
+hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2
+          differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux
+
+
+
+(* TODO: differentiable etc. *)
+
+
+subsection \<open>Continuity\<close>
+
+lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma"
+  by (rule DERIV_continuous_on has_field_derivative_rGamma)+
+
+lemma continuous_on_Gamma [continuous_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A Gamma"
+  by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast
+
+lemma isCont_rGamma [continuous_intros]:
+  "isCont f z \<Longrightarrow> isCont (\<lambda>x. rGamma (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_rGamma]])
+
+lemma isCont_Gamma [continuous_intros]:
+  "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Gamma]])
+
+
+
+text \<open>The complex Gamma function\<close>
+
+instantiation complex :: Gamma
+begin
+
+definition rGamma_complex :: "complex \<Rightarrow> complex" where
+  "rGamma_complex z = lim (rGamma_series z)"
+
+lemma rGamma_series_complex_converges:
+        "convergent (rGamma_series (z :: complex))" (is "?thesis1")
+  and rGamma_complex_altdef:
+        "rGamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2")
+proof -
+  have "?thesis1 \<and> ?thesis2"
+  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+    case False
+    have "rGamma_series z \<longlonglongrightarrow> exp (- ln_Gamma z)"
+    proof (rule Lim_transform_eventually)
+      from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE)
+      from this(1) uniformly_convergent_imp_convergent[OF this(2), of z]
+        have "ln_Gamma_series z \<longlonglongrightarrow> lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff)
+      thus "(\<lambda>n. exp (-ln_Gamma_series z n)) \<longlonglongrightarrow> exp (- ln_Gamma z)"
+        unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus)
+      from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False
+        show "eventually (\<lambda>n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially"
+        by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def)
+    qed
+    with False show ?thesis
+      by (auto simp: convergent_def rGamma_complex_def intro!: limI)
+  next
+    case True
+    then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases')
+    also have "rGamma_series \<dots> \<longlonglongrightarrow> 0"
+      by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const)
+    finally show ?thesis using True
+      by (auto simp: rGamma_complex_def convergent_def intro!: limI)
+  qed
+  thus "?thesis1" "?thesis2" by blast+
+qed
+
+context
+begin
+
+(* TODO: duplication *)
+private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)"
+proof -
+  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
+  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
+             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
+      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
+    also from n have "\<dots> = ?f n * rGamma_series z n"
+      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
+    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
+  qed
+  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
+    using rGamma_series_complex_converges
+    by (intro tendsto_intros lim_inverse_n)
+       (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def)
+  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
+  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
+    by (rule Lim_transform_eventually)
+  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
+    using rGamma_series_complex_converges
+    by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff)
+  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
+qed
+
+private lemma has_field_derivative_rGamma_complex_no_nonpos_Int:
+  assumes "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
+proof -
+  have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z
+  proof (subst DERIV_cong_ev[OF refl _ refl])
+    from that have "eventually (\<lambda>t. t \<in> ball z (Re z/2)) (nhds z)"
+      by (intro eventually_nhds_in_nhd) simp_all
+    thus "eventually (\<lambda>t. rGamma t = exp (- ln_Gamma t)) (nhds z)"
+      using no_nonpos_Int_in_ball_complex[OF that]
+      by (auto elim!: eventually_mono simp: rGamma_complex_altdef)
+  next
+    have "z \<notin> \<real>\<^sub>\<le>\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff)
+    with that show "((\<lambda>t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)"
+     by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef)
+  qed
+
+  from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
+  proof (induction "nat \<lfloor>1 - Re z\<rfloor>" arbitrary: z)
+    case (Suc n z)
+    from Suc.prems have z: "z \<noteq> 0" by auto
+    from Suc.hyps have "n = nat \<lfloor>- Re z\<rfloor>" by linarith
+    hence A: "n = nat \<lfloor>1 - Re (z + 1)\<rfloor>" by simp
+    from Suc.prems have B: "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp)
+
+    have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1)) z) has_field_derivative
+      -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)"
+      by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps)
+    also have "(\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) = rGamma"
+      by (simp add: rGamma_complex_plus1)
+    also from z have "Digamma (z + 1) * z - 1 = z * Digamma z"
+      by (subst Digamma_plus1) (simp_all add: field_simps)
+    also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z"
+      by (simp add: rGamma_complex_plus1[of z, symmetric])
+    finally show ?case .
+  qed (intro diff, simp)
+qed
+
+private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1"
+proof -
+  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
+                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
+  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
+  thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI)
+qed
+
+private lemma has_field_derivative_rGamma_complex_nonpos_Int:
+  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))"
+proof (induction n)
+  case 0
+  have A: "(0::complex) + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by simp
+  have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)"
+    by (rule derivative_eq_intros DERIV_chain refl
+             has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1)
+    thus ?case by (simp add: rGamma_complex_plus1)
+next
+  case (Suc n)
+  hence A: "(rGamma has_field_derivative (-1)^n * fact n)
+                (at (- of_nat (Suc n) + 1 :: complex))" by simp
+   have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative
+             (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))"
+     by (rule derivative_eq_intros refl A DERIV_chain)+
+        (simp add: algebra_simps rGamma_complex_altdef)
+  thus ?case by (simp add: rGamma_complex_plus1)
+qed
+
+instance proof
+  fix z :: complex show "(rGamma z = 0) \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
+    by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases')
+next
+  fix z :: complex assume "\<And>n. z \<noteq> - of_nat n"
+  hence "z \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases')
+  from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this]
+    show "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
+                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma z +
+              rGamma z * d * (y - z)) /\<^sub>R  cmod (y - z)) \<midarrow>z\<rightarrow> 0"
+      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
+                    netlimit_at of_real_def[symmetric] suminf_def)
+next
+  fix n :: nat
+  from has_field_derivative_rGamma_complex_nonpos_Int[of n]
+  show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} *
+                  (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"
+    by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def)
+next
+  fix z :: complex
+  from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z"
+    by (simp add: convergent_LIMSEQ_iff rGamma_complex_def)
+  thus "let fact' = \<lambda>n. setprod of_nat {1..n};
+            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
+            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
+        in  (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z"
+    by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
+                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
+qed
+
+end
+end
+
+
+lemma Gamma_complex_altdef:
+  "Gamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))"
+  unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus)
+
+lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)"
+proof -
+  have "rGamma_series (cnj z) = (\<lambda>n. cnj (rGamma_series z n))"
+    by (intro ext) (simp_all add: rGamma_series_def exp_cnj)
+  also have "... \<longlonglongrightarrow> cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros)
+  finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI])
+qed
+
+lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)"
+  unfolding Gamma_def by (simp add: cnj_rGamma)
+
+lemma Gamma_complex_real:
+  "z \<in> \<real> \<Longrightarrow> Gamma z \<in> (\<real> :: complex set)" and rGamma_complex_real: "z \<in> \<real> \<Longrightarrow> rGamma z \<in> \<real>"
+  by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma)
+
+lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)"
+  using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast
+
+lemma holomorphic_on_rGamma: "rGamma holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma)
+
+lemma analytic_on_rGamma: "rGamma analytic_on A"
+  unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma)
+
+
+lemma field_differentiable_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma field_differentiable (at z within A)"
+  using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto
+
+lemma holomorphic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma)
+
+lemma analytic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma analytic_on A"
+  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
+     (auto intro!: holomorphic_on_Gamma)
+
+lemma has_field_derivative_rGamma_complex' [derivative_intros]:
+  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-Re z\<rfloor>) * fact (nat \<lfloor>-Re z\<rfloor>) else
+        -rGamma z * Digamma z)) (at z within A)"
+  using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros]
+
+
+lemma field_differentiable_Polygamma:
+  fixes z::complex
+  shows
+  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Polygamma n field_differentiable (at z within A)"
+  using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto
+
+lemma holomorphic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma)
+
+lemma analytic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n analytic_on A"
+  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
+     (auto intro!: holomorphic_on_Polygamma)
+
+
+
+text \<open>The real Gamma function\<close>
+
+lemma rGamma_series_real:
+  "eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially"
+  using eventually_gt_at_top[of "0 :: nat"]
+proof eventually_elim
+  fix n :: nat assume n: "n > 0"
+  have "Re (rGamma_series (of_real x) n) =
+          Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))"
+    using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real)
+  also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) /
+                              (fact n * (exp (x * ln (real_of_nat n))))))"
+    by (subst exp_of_real) simp
+  also from n have "\<dots> = rGamma_series x n"
+    by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def)
+  finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" ..
+qed
+
+instantiation real :: Gamma
+begin
+
+definition "rGamma_real x = Re (rGamma (of_real x :: complex))"
+
+instance proof
+  fix x :: real
+  have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def)
+  also have "of_real \<dots> = rGamma (of_real x :: complex)"
+    by (intro of_real_Re rGamma_complex_real) simp_all
+  also have "\<dots> = 0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff)
+  also have "\<dots> \<longleftrightarrow> (\<exists>n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases')
+  finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp
+next
+  fix x :: real assume "\<And>n. x \<noteq> - of_nat n"
+  hence x: "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')
+  then have "x \<noteq> 0" by auto
+  with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
+    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
+                  simp: Polygamma_of_real rGamma_real_def [abs_def])
+  thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k))
+                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma x +
+              rGamma x * d * (y - x)) /\<^sub>R  norm (y - x)) \<midarrow>x\<rightarrow> 0"
+      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
+                    netlimit_at of_real_def[symmetric] suminf_def)
+next
+  fix n :: nat
+  have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))"
+    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
+                  simp: Polygamma_of_real rGamma_real_def [abs_def])
+  thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} *
+                  (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"
+    by (simp add: has_field_derivative_def has_derivative_def fact_setprod netlimit_at Let_def)
+next
+  fix x :: real
+  have "rGamma_series x \<longlonglongrightarrow> rGamma x"
+  proof (rule Lim_transform_eventually)
+    show "(\<lambda>n. Re (rGamma_series (of_real x) n)) \<longlonglongrightarrow> rGamma x" unfolding rGamma_real_def
+      by (intro tendsto_intros)
+  qed (insert rGamma_series_real, simp add: eq_commute)
+  thus "let fact' = \<lambda>n. setprod of_nat {1..n};
+            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
+            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
+        in  (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x"
+    by (simp add: fact_setprod pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
+                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
+qed
+
+end
+
+
+lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)"
+  unfolding rGamma_real_def using rGamma_complex_real by simp
+
+lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)"
+  unfolding Gamma_def by (simp add: rGamma_complex_of_real)
+
+lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))"
+  by (rule sym, rule limI, rule tendsto_intros)
+
+lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))"
+  by (rule sym, rule limI, rule tendsto_intros)
+
+lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))"
+  using rGamma_complex_real[OF Reals_of_real[of x]]
+  by (elim Reals_cases)
+     (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real)
+
+lemma ln_Gamma_series_complex_of_real:
+  "x > 0 \<Longrightarrow> n > 0 \<Longrightarrow> ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)"
+proof -
+  assume xn: "x > 0" "n > 0"
+  have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k
+    using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps)
+  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real)
+qed
+
+lemma ln_Gamma_real_converges:
+  assumes "(x::real) > 0"
+  shows   "convergent (ln_Gamma_series x)"
+proof -
+  have "(\<lambda>n. ln_Gamma_series (complex_of_real x) n) \<longlonglongrightarrow> ln_Gamma (of_real x)" using assms
+    by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff)
+  moreover from eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. complex_of_real (ln_Gamma_series x n) =
+            ln_Gamma_series (complex_of_real x) n) sequentially"
+    by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms)
+  ultimately have "(\<lambda>n. complex_of_real (ln_Gamma_series x n)) \<longlonglongrightarrow> ln_Gamma (of_real x)"
+    by (subst tendsto_cong) assumption+
+  from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def)
+qed
+
+lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \<Longrightarrow> ln_Gamma_series x \<longlonglongrightarrow> ln_Gamma x"
+  using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff)
+
+lemma ln_Gamma_complex_of_real: "x > 0 \<Longrightarrow> ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)"
+proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually)
+  assume x: "x > 0"
+  show "eventually (\<lambda>n. of_real (ln_Gamma_series x n) =
+            ln_Gamma_series (complex_of_real x) n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by eventually_elim (simp add: ln_Gamma_series_complex_of_real x)
+qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def)
+
+lemma Gamma_real_pos_exp: "x > (0 :: real) \<Longrightarrow> Gamma x = exp (ln_Gamma x)"
+  by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff
+                 ln_Gamma_complex_of_real exp_of_real)
+
+lemma ln_Gamma_real_pos: "x > 0 \<Longrightarrow> ln_Gamma x = ln (Gamma x :: real)"
+  unfolding Gamma_real_pos_exp by simp
+
+lemma Gamma_real_pos: "x > (0::real) \<Longrightarrow> Gamma x > 0"
+  by (simp add: Gamma_real_pos_exp)
+
+lemma has_field_derivative_ln_Gamma_real [derivative_intros]:
+  assumes x: "x > (0::real)"
+  shows "(ln_Gamma has_field_derivative Digamma x) (at x)"
+proof (subst DERIV_cong_ev[OF refl _ refl])
+  from assms show "((Re \<circ> ln_Gamma \<circ> complex_of_real) has_field_derivative Digamma x) (at x)"
+    by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex
+             simp: Polygamma_of_real o_def)
+  from eventually_nhds_in_nhd[of x "{0<..}"] assms
+    show "eventually (\<lambda>y. ln_Gamma y = (Re \<circ> ln_Gamma \<circ> of_real) y) (nhds x)"
+    by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open)
+qed
+
+declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros]
+
+
+lemma has_field_derivative_rGamma_real' [derivative_intros]:
+  "(rGamma has_field_derivative (if x \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-x\<rfloor>) * fact (nat \<lfloor>-x\<rfloor>) else
+        -rGamma x * Digamma x)) (at x within A)"
+  using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros]
+
+lemma Polygamma_real_odd_pos:
+  assumes "(x::real) \<notin> \<int>\<^sub>\<le>\<^sub>0" "odd n"
+  shows   "Polygamma n x > 0"
+proof -
+  from assms have "x \<noteq> 0" by auto
+  with assms show ?thesis
+    unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
+    by (auto simp: zero_less_power_eq simp del: power_Suc
+             dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos)
+qed
+
+lemma Polygamma_real_even_neg:
+  assumes "(x::real) > 0" "n > 0" "even n"
+  shows   "Polygamma n x < 0"
+  using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
+  by (auto intro!: mult_pos_pos suminf_pos)
+
+lemma Polygamma_real_strict_mono:
+  assumes "x > 0" "x < (y::real)" "even n"
+  shows   "Polygamma n x < Polygamma n y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> > 0"
+    by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases)
+  finally show ?thesis by simp
+qed
+
+lemma Polygamma_real_strict_antimono:
+  assumes "x > 0" "x < (y::real)" "odd n"
+  shows   "Polygamma n x > Polygamma n y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> < 0"
+    by (intro mult_pos_neg Polygamma_real_even_neg) simp_all
+  finally show ?thesis by simp
+qed
+
+lemma Polygamma_real_mono:
+  assumes "x > 0" "x \<le> (y::real)" "even n"
+  shows   "Polygamma n x \<le> Polygamma n y"
+  using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2)
+  by (cases "x = y") simp_all
+
+lemma Digamma_real_ge_three_halves_pos:
+  assumes "x \<ge> 3/2"
+  shows   "Digamma (x :: real) > 0"
+proof -
+  have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos)
+  also from assms have "\<dots> \<le> Digamma x" by (intro Polygamma_real_mono) simp_all
+  finally show ?thesis .
+qed
+
+lemma ln_Gamma_real_strict_mono:
+  assumes "x \<ge> 3/2" "x < y"
+  shows   "ln_Gamma (x :: real) < ln_Gamma y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> ln_Gamma y - ln_Gamma x = (y - x) * Digamma \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Digamma \<xi> > 0"
+    by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all
+  finally show ?thesis by simp
+qed
+
+lemma Gamma_real_strict_mono:
+  assumes "x \<ge> 3/2" "x < y"
+  shows   "Gamma (x :: real) < Gamma y"
+proof -
+  from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp
+  also have "\<dots> < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms)
+  also from Gamma_real_pos_exp[of y] assms have "\<dots> = Gamma y" by simp
+  finally show ?thesis .
+qed
+
+lemma log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)"
+  by (rule convex_on_realI[of _ _ Digamma])
+     (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos
+           simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases')
+
+
+subsection \<open>Beta function\<close>
+
+definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)"
+
+lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)"
+  by (simp add: inverse_eq_divide Beta_def Gamma_def)
+
+lemma Beta_commute: "Beta a b = Beta b a"
+  unfolding Beta_def by (simp add: ac_simps)
+
+lemma has_field_derivative_Beta1 [derivative_intros]:
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "((\<lambda>x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y))))
+               (at x within A)" unfolding Beta_altdef
+  by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps)
+
+lemma Beta_pole1: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
+  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
+
+lemma Beta_pole2: "y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
+  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
+
+lemma Beta_zero: "x + y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
+  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
+
+lemma has_field_derivative_Beta2 [derivative_intros]:
+  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "((\<lambda>y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y))))
+               (at y within A)"
+  using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac)
+
+lemma Beta_plus1_plus1:
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "Beta (x + 1) y + Beta x (y + 1) = Beta x y"
+proof -
+  have "Beta (x + 1) y + Beta x (y + 1) =
+            (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)"
+    by (simp add: Beta_altdef add_divide_distrib algebra_simps)
+  also have "\<dots> = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))"
+    by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps)
+  also from assms have "\<dots> = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp
+  finally show ?thesis .
+qed
+
+lemma Beta_plus1_left:
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(x + y) * Beta (x + 1) y = x * Beta x y"
+proof -
+  have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))"
+    unfolding Beta_altdef by (simp only: ac_simps)
+  also have "\<dots> = x * Beta x y" unfolding Beta_altdef
+     by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps)
+  finally show ?thesis .
+qed
+
+lemma Beta_plus1_right:
+  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(x + y) * Beta x (y + 1) = y * Beta x y"
+  using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute)
+
+lemma Gamma_Gamma_Beta:
+  assumes "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "Gamma x * Gamma y = Beta x y * Gamma (x + y)"
+  unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"]
+  by (simp add: rGamma_inverse_Gamma)
+
+
+
+subsection \<open>Legendre duplication theorem\<close>
+
+context
+begin
+
+private lemma Gamma_legendre_duplication_aux:
+  fixes z :: "'a :: Gamma"
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)"
+proof -
+  let ?powr = "\<lambda>b a. exp (a * of_real (ln (of_nat b)))"
+  let ?h = "\<lambda>n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) *
+                exp (1/2 * of_real (ln (real_of_nat n)))"
+  {
+    fix z :: 'a assume z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    let ?g = "\<lambda>n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n /
+                      Gamma_series' (2*z) (2*n)"
+    have "eventually (\<lambda>n. ?g n = ?h n) sequentially" using eventually_gt_at_top
+    proof eventually_elim
+      fix n :: nat assume n: "n > 0"
+      let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a"
+      have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp
+      have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) /
+                (pochhammer z n * pochhammer (z + 1/2) n)"
+        by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac)
+      have B: "Gamma_series' (2*z) (2*n) =
+                       ?f' * ?powr 2 (2*z) * ?powr n (2*z) /
+                       (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n
+        by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double)
+      from z have "pochhammer z n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
+      moreover from z have "pochhammer (z + 1/2) n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
+      ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) =
+         ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))"
+        using n unfolding A B by (simp add: divide_simps exp_minus)
+      also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)"
+        by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib)
+      finally show "?g n = ?h n" by (simp only: mult_ac)
+    qed
+
+    moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
+    hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
+      using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"]
+      by (intro tendsto_intros Gamma_series'_LIMSEQ)
+         (simp_all add: o_def subseq_def Gamma_eq_zero_iff)
+    ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
+      by (rule Lim_transform_eventually)
+  } note lim = this
+
+  from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
+  from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all
+  with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real)
+  from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis
+    by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
+qed
+
+(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is
+   infinitely differentiable *)
+private lemma Gamma_reflection_aux:
+  defines "h \<equiv> \<lambda>z::complex. if z \<in> \<int> then 0 else
+                 (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
+  defines "a \<equiv> complex_of_real pi"
+  obtains h' where "continuous_on UNIV h'" "\<And>z. (h has_field_derivative (h' z)) (at z)"
+proof -
+  define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n
+  define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z
+  define g where "g n = complex_of_real (sin_coeff (n+1))" for n
+  define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z
+  have a_nz: "a \<noteq> 0" unfolding a_def by simp
+
+  have "(\<lambda>n. f n * (a*z)^n) sums (F z) \<and> (\<lambda>n. g n * (a*z)^n) sums (G z)"
+    if "abs (Re z) < 1" for z
+  proof (cases "z = 0"; rule conjI)
+    assume "z \<noteq> 0"
+    note z = this that
+
+    from z have sin_nz: "sin (a*z) \<noteq> 0" unfolding a_def by (auto simp: sin_eq_0)
+    have "(\<lambda>n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"]
+      by (simp add: scaleR_conv_of_real)
+    from sums_split_initial_segment[OF this, of 1]
+      have "(\<lambda>n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac)
+    from sums_mult[OF this, of "inverse (a*z)"] z a_nz
+      have A: "(\<lambda>n. g n * (a*z)^n) sums (sin (a*z)/(a*z))"
+      by (simp add: field_simps g_def)
+    with z show "(\<lambda>n. g n * (a*z)^n) sums (G z)" by (simp add: G_def)
+    from A z a_nz sin_nz have g_nz: "(\<Sum>n. g n * (a*z)^n) \<noteq> 0" by (simp add: sums_iff g_def)
+
+    have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def)
+    from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1]
+    have "(\<lambda>n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))"
+      by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def)
+    from sums_mult[OF this, of "inverse z"] z assms
+      show "(\<lambda>n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def)
+  next
+    assume z: "z = 0"
+    have "(\<lambda>n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp
+    with z show "(\<lambda>n. f n * (a * z) ^ n) sums (F z)"
+      by (simp add: f_def F_def sin_coeff_def cos_coeff_def)
+    have "(\<lambda>n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp
+    with z show "(\<lambda>n. g n * (a * z) ^ n) sums (G z)"
+      by (simp add: g_def G_def sin_coeff_def cos_coeff_def)
+  qed
+  note sums = conjunct1[OF this] conjunct2[OF this]
+
+  define h2 where [abs_def]:
+    "h2 z = (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z
+  define POWSER where [abs_def]: "POWSER f z = (\<Sum>n. f n * (z^n :: complex))" for f z
+  define POWSER' where [abs_def]: "POWSER' f z = (\<Sum>n. diffs f n * (z^n))" for f and z :: complex
+  define h2' where [abs_def]:
+    "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) /
+      (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z
+
+  have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t
+  proof -
+    from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases simp: dist_0_norm)
+    hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)"
+      unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def)
+    also have "a*cot (a*t) - 1/t = (F t) / (G t)"
+      using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def)
+    also have "\<dots> = (\<Sum>n. f n * (a*t)^n) / (\<Sum>n. g n * (a*t)^n)"
+      using sums[of t] that by (simp add: sums_iff dist_0_norm)
+    finally show "h t = h2 t" by (simp only: h2_def)
+  qed
+
+  let ?A = "{z. abs (Re z) < 1}"
+  have "open ({z. Re z < 1} \<inter> {z. Re z > -1})"
+    using open_halfspace_Re_gt open_halfspace_Re_lt by auto
+  also have "({z. Re z < 1} \<inter> {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto
+  finally have open_A: "open ?A" .
+  hence [simp]: "interior ?A = ?A" by (simp add: interior_open)
+
+  have summable_f: "summable (\<lambda>n. f n * z^n)" for z
+    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
+       (simp_all add: norm_mult a_def del: of_real_add)
+  have summable_g: "summable (\<lambda>n. g n * z^n)" for z
+    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
+       (simp_all add: norm_mult a_def del: of_real_add)
+  have summable_fg': "summable (\<lambda>n. diffs f n * z^n)" "summable (\<lambda>n. diffs g n * z^n)" for z
+    by (intro termdiff_converges_all summable_f summable_g)+
+  have "(POWSER f has_field_derivative (POWSER' f z)) (at z)"
+               "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z
+    unfolding POWSER_def POWSER'_def
+    by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+
+  note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def]
+  have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z"
+    for z unfolding POWSER_def POWSER'_def
+    by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+
+  note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def]
+
+  {
+    fix z :: complex assume z: "abs (Re z) < 1"
+    define d where "d = \<i> * of_real (norm z + 1)"
+    have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add)
+    have "eventually (\<lambda>z. h z = h2 z) (nhds z)"
+      using eventually_nhds_in_nhd[of z ?A] using h_eq z
+      by (auto elim!: eventually_mono simp: dist_0_norm)
+
+    moreover from sums(2)[OF z] z have nz: "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
+      unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def)
+    have A: "z \<in> \<int> \<longleftrightarrow> z = 0" using z by (auto elim!: Ints_cases)
+    have no_int: "1 + z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of "1+z" 1] A
+      by (auto elim!: nonpos_Ints_cases)
+    have no_int': "1 - z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of 1 "1-z"] A
+      by (auto elim!: nonpos_Ints_cases)
+    from no_int no_int' have no_int: "1 - z \<notin> \<int>\<^sub>\<le>\<^sub>0" "1 + z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
+    have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def
+      by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+)
+         (auto simp: h2'_def POWSER_def field_simps power2_eq_square)
+    ultimately have deriv: "(h has_field_derivative h2' z) (at z)"
+      by (subst DERIV_cong_ev[OF refl _ refl])
+
+    from sums(2)[OF z] z have "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
+      unfolding G_def by (auto simp: sums_iff a_def sin_eq_0)
+    hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def
+      by (intro continuous_intros cont
+            continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto
+    note deriv and this
+  } note A = this
+
+  interpret h: periodic_fun_simple' h
+  proof
+    fix z :: complex
+    show "h (z + 1) = h z"
+    proof (cases "z \<in> \<int>")
+      assume z: "z \<notin> \<int>"
+      hence A: "z + 1 \<notin> \<int>" "z \<noteq> 0" using Ints_diff[of "z+1" 1] by auto
+      hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)"
+        by (subst (1 2) Digamma_plus1) simp_all
+      with A z show "h (z + 1) = h z"
+        by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def)
+    qed (simp add: h_def)
+  qed
+
+  have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z
+  proof -
+    have "((\<lambda>z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)"
+      by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
+         (insert z, auto intro!: derivative_eq_intros)
+    hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1)
+    moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all
+    ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique)
+  qed
+
+  define h2'' where "h2'' z = h2' (z - of_int \<lfloor>Re z\<rfloor>)" for z
+  have deriv: "(h has_field_derivative h2'' z) (at z)" for z
+  proof -
+    fix z :: complex
+    have B: "\<bar>Re z - real_of_int \<lfloor>Re z\<rfloor>\<bar> < 1" by linarith
+    have "((\<lambda>t. h (t - of_int \<lfloor>Re z\<rfloor>)) has_field_derivative h2'' z) (at z)"
+      unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
+                            (insert B, auto intro!: derivative_intros)
+    thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int)
+  qed
+
+  have cont: "continuous_on UNIV h2''"
+  proof (intro continuous_at_imp_continuous_on ballI)
+    fix z :: complex
+    define r where "r = \<lfloor>Re z\<rfloor>"
+    define A where "A = {t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}"
+    have "continuous_on A (\<lambda>t. h2' (t - of_int r))" unfolding A_def
+      by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros)
+         (simp_all add: abs_real_def)
+    moreover have "h2'' t = h2' (t - of_int r)" if t: "t \<in> A" for t
+    proof (cases "Re t \<ge> of_int r")
+      case True
+      from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def)
+      with True have "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor>" unfolding r_def by linarith
+      thus ?thesis by (auto simp: r_def h2''_def)
+    next
+      case False
+      from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def)
+      with False have t': "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor> - 1" unfolding r_def by linarith
+      moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)"
+        by (intro h2'_eq) simp_all
+      ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t')
+    qed
+    ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl])
+    moreover {
+      have "open ({t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t})"
+        by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt)
+      also have "{t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t} = A"
+        unfolding A_def by blast
+      finally have "open A" .
+    }
+    ultimately have C: "isCont h2'' t" if "t \<in> A" for t using that
+      by (subst (asm) continuous_on_eq_continuous_at) auto
+    have "of_int r - 1 < Re z" "Re z  < of_int r + 1" unfolding r_def by linarith+
+    thus "isCont h2'' z" by (intro C) (simp_all add: A_def)
+  qed
+
+  from that[OF cont deriv] show ?thesis .
+qed
+
+lemma Gamma_reflection_complex:
+  fixes z :: complex
+  shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)"
+proof -
+  let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"
+  define g where [abs_def]: "g z = (if z \<in> \<int> then of_real pi else ?g z)" for z :: complex
+  let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
+  define h where [abs_def]: "h z = (if z \<in> \<int> then 0 else ?h z)" for z :: complex
+
+  \<comment> \<open>@{term g} is periodic with period 1.\<close>
+  interpret g: periodic_fun_simple' g
+  proof
+    fix z :: complex
+    show "g (z + 1) = g z"
+    proof (cases "z \<in> \<int>")
+      case False
+      hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def)
+      also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)"
+        using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints
+        by (subst Beta_plus1_left [symmetric]) auto
+      also have "\<dots> * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))"
+        using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints
+        by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi)
+      also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)"
+        using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def)
+      finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto
+    qed (simp add: g_def)
+  qed
+
+  \<comment> \<open>@{term g} is entire.\<close>
+  have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex
+  proof (cases "z \<in> \<int>")
+    let ?h' = "\<lambda>z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) +
+                     of_real pi * cos (z * of_real pi))"
+    case False
+    from False have "eventually (\<lambda>t. t \<in> UNIV - \<int>) (nhds z)"
+      by (intro eventually_nhds_in_open) (auto simp: open_Diff)
+    hence "eventually (\<lambda>t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def)
+    moreover {
+      from False Ints_diff[of 1 "1-z"] have "1 - z \<notin> \<int>" by auto
+      hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints
+        by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def)
+      also from False have "sin (of_real pi * z) \<noteq> 0" by (subst sin_eq_0) auto
+      hence "?h' z = h z * g z"
+        using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def)
+      finally have "(?g has_field_derivative (h z * g z)) (at z)" .
+    }
+    ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
+  next
+    case True
+    then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases)
+    let ?t = "(\<lambda>z::complex. if z = 0 then 1 else sin z / z) \<circ> (\<lambda>z. of_real pi * z)"
+    have deriv_0: "(g has_field_derivative 0) (at 0)"
+    proof (subst DERIV_cong_ev[OF refl _ refl])
+      show "eventually (\<lambda>z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)"
+        using eventually_nhds_ball[OF zero_less_one, of "0::complex"]
+      proof eventually_elim
+        fix z :: complex assume z: "z \<in> ball 0 1"
+        show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z"
+        proof (cases "z = 0")
+          assume z': "z \<noteq> 0"
+          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm)
+          from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp
+          with z'' z' show ?thesis by (simp add: g_def ac_simps)
+        qed (simp add: g_def)
+      qed
+      have "(?t has_field_derivative (0 * of_real pi)) (at 0)"
+        using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"]
+        by (intro DERIV_chain) simp_all
+      thus "((\<lambda>z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)"
+        by (auto intro!: derivative_eq_intros simp: o_def)
+    qed
+
+    have "((g \<circ> (\<lambda>x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))"
+      using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros)
+    also have "g \<circ> (\<lambda>x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int)
+    finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def)
+  qed
+
+  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z
+  proof (cases "z \<in> \<int>")
+    case True
+    with that have "z = 0 \<or> z = 1" by (force elim!: Ints_cases)
+    moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0"
+      using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square)
+    moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1"
+        using fraction_not_in_ints[where 'a = complex, of 2 1]
+        by (simp add: g_def power2_eq_square Beta_def algebra_simps)
+    ultimately show ?thesis by force
+  next
+    case False
+    hence z: "z/2 \<notin> \<int>" "(z+1)/2 \<notin> \<int>" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases)
+    hence z': "z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "(z+1)/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
+    from z have "1-z/2 \<notin> \<int>" "1-((z+1)/2) \<notin> \<int>"
+      using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto
+    hence z'': "1-z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "1-((z+1)/2) \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
+    from z have "g (z/2) * g ((z+1)/2) =
+      (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) *
+      (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))"
+      by (simp add: g_def)
+    also from z' Gamma_legendre_duplication_aux[of "z/2"]
+      have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z"
+      by (simp add: add_divide_distrib)
+    also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"]
+      have "Gamma (1-z/2) * Gamma (1-(z+1)/2) =
+              Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))"
+      by (simp add: add_divide_distrib ac_simps)
+    finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) *
+                    (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))"
+      by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real)
+    also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)"
+      using cos_sin_eq[of "- of_real pi * z/2", symmetric]
+      by (simp add: ring_distribs add_divide_distrib ac_simps)
+    also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)"
+      by (subst sin_times_cos) (simp add: field_simps)
+    also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z"
+      using \<open>z \<notin> \<int>\<close> by (simp add: g_def)
+    finally show ?thesis .
+  qed
+  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z
+  proof -
+    define r where "r = \<lfloor>Re z / 2\<rfloor>"
+    have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int)
+    also have "of_int (2*r) = 2 * of_int r" by simp
+    also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+
+    hence "Gamma (1/2)^2 * g (z - 2 * of_int r) =
+                   g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)"
+      unfolding r_def by (intro g_eq[symmetric]) simp_all
+    also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp
+    also have "g \<dots> = g (z/2)" by (rule g.minus_of_int)
+    also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp
+    also have "g \<dots> = g ((z+1)/2)" by (rule g.minus_of_int)
+    finally show ?thesis ..
+  qed
+
+  have g_nz [simp]: "g z \<noteq> 0" for z :: complex
+  unfolding g_def using Ints_diff[of 1 "1 - z"]
+    by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int)
+
+  have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z
+  proof -
+    have "((\<lambda>t. g (t/2) * g ((t+1)/2)) has_field_derivative
+                       (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
+      by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps)
+    hence "((\<lambda>t. Gamma (1/2)^2 * g t) has_field_derivative
+              Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
+      by (subst (1 2) g_eq[symmetric]) simp
+    from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"]
+      have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)"
+      using fraction_not_in_ints[where 'a = complex, of 2 1]
+      by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints)
+    moreover have "(g has_field_derivative (g z * h z)) (at z)"
+      using g_g'[of z] by (simp add: ac_simps)
+    ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)"
+      by (intro DERIV_unique)
+    thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp
+  qed
+
+  obtain h' where h'_cont: "continuous_on UNIV h'" and
+                  h_h': "\<And>z. (h has_field_derivative h' z) (at z)"
+     unfolding h_def by (erule Gamma_reflection_aux)
+
+  have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z
+  proof -
+    have "((\<lambda>t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative
+                       ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)"
+      by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2])
+    hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)"
+      by (subst (asm) h_eq[symmetric])
+    from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique)
+  qed
+
+  have h'_zero: "h' z = 0" for z
+  proof -
+    define m where "m = max 1 \<bar>Re z\<bar>"
+    define B where "B = {t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
+    have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter>
+                  {t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})"
+      (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le
+                                 closed_halfspace_Im_ge closed_halfspace_Im_le)
+    also have "?B = B" unfolding B_def by fastforce
+    finally have "closed B" .
+    moreover have "bounded B" unfolding bounded_iff
+    proof (intro ballI exI)
+      fix t assume t: "t \<in> B"
+      have "norm t \<le> \<bar>Re t\<bar> + \<bar>Im t\<bar>" by (rule cmod_le)
+      also from t have "\<bar>Re t\<bar> \<le> m" unfolding B_def by blast
+      also from t have "\<bar>Im t\<bar> \<le> \<bar>Im z\<bar>" unfolding B_def by blast
+      finally show "norm t \<le> m + \<bar>Im z\<bar>" by - simp
+    qed
+    ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast
+
+    define M where "M = (SUP z:B. norm (h' z))"
+    have "compact (h' ` B)"
+      by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+
+    hence bdd: "bdd_above ((\<lambda>z. norm (h' z)) ` B)"
+      using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded)
+    have "norm (h' z) \<le> M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def)
+    also have "M \<le> M/2"
+    proof (subst M_def, subst cSUP_le_iff)
+      have "z \<in> B" unfolding B_def m_def by simp
+      thus "B \<noteq> {}" by auto
+    next
+      show "\<forall>z\<in>B. norm (h' z) \<le> M/2"
+      proof
+        fix t :: complex assume t: "t \<in> B"
+        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm)
+        also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp
+        also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))"
+          by (rule norm_triangle_ineq)
+        also from t have "abs (Re ((t + 1)/2)) \<le> m" unfolding m_def B_def by auto
+        with t have "t/2 \<in> B" "(t+1)/2 \<in> B" unfolding B_def by auto
+        hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \<le> M + M" unfolding M_def
+          by (intro add_mono cSUP_upper bdd) (auto simp: B_def)
+        also have "(M + M) / 4 = M / 2" by simp
+        finally show "norm (h' t) \<le> M/2" by - simp_all
+      qed
+    qed (insert bdd, auto simp: cball_eq_empty)
+    hence "M \<le> 0" by simp
+    finally show "h' z = 0" by simp
+  qed
+  have h_h'_2: "(h has_field_derivative 0) (at z)" for z
+    using h_h'[of z] h'_zero[of z] by simp
+
+  have g_real: "g z \<in> \<real>" if "z \<in> \<real>" for z
+    unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real)
+  have h_real: "h z \<in> \<real>" if "z \<in> \<real>" for z
+    unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real)
+  have g_nz: "g z \<noteq> 0" for z unfolding g_def using Ints_diff[of 1 "1-z"]
+    by (auto simp: Gamma_eq_zero_iff sin_eq_0)
+
+  from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c"
+    by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm)
+  then obtain c where c: "\<And>z. h z = c" by auto
+  have "\<exists>u. u \<in> closed_segment 0 1 \<and> Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))"
+    by (intro complex_mvt_line g_g')
+    find_theorems name:deriv Reals
+  then guess u by (elim exE conjE) note u = this
+  from u(1) have u': "u \<in> \<real>" unfolding closed_segment_def
+    by (auto simp: scaleR_conv_of_real)
+  from u' g_real[of u] g_nz[of u] have "Re (g u) \<noteq> 0" by (auto elim!: Reals_cases)
+  with u(2) c[of u] g_real[of u] g_nz[of u] u'
+    have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1)
+  with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases)
+  with c have A: "h z * g z = 0" for z by simp
+  hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
+  hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
+  then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
+  from this[of 0] have "c' = pi" unfolding g_def by simp
+  with c have "g z = pi" by simp
+
+  show ?thesis
+  proof (cases "z \<in> \<int>")
+    case False
+    with \<open>g z = pi\<close> show ?thesis by (auto simp: g_def divide_simps)
+  next
+    case True
+    then obtain n where n: "z = of_int n" by (elim Ints_cases)
+    with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force
+    moreover have "of_int (1 - n) \<in> \<int>\<^sub>\<le>\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp
+    ultimately show ?thesis using n
+      by (cases "n \<le> 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int)
+  qed
+qed
+
+lemma rGamma_reflection_complex:
+  "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi"
+  using Gamma_reflection_complex[of z]
+    by (simp add: Gamma_def divide_simps split: if_split_asm)
+
+lemma rGamma_reflection_complex':
+  "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi"
+proof -
+  have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))"
+    using rGamma_plus1[of "-z", symmetric] by simp
+  also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi"
+    by (rule rGamma_reflection_complex)
+  finally show ?thesis by simp
+qed
+
+lemma Gamma_reflection_complex':
+  "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))"
+  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac)
+
+
+
+lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi"
+proof -
+  from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1]
+    have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square)
+  hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp
+  also have "\<dots> = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all
+  finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all
+  moreover have "Gamma (1/2 :: real) \<ge> 0" using Gamma_real_pos[of "1/2"] by simp
+  ultimately show ?thesis by (rule real_sqrt_unique [symmetric])
+qed
+
+lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)"
+proof -
+  have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp
+  also have "\<dots> = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real)
+  finally show ?thesis .
+qed
+
+lemma Gamma_legendre_duplication:
+  fixes z :: complex
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "Gamma z * Gamma (z + 1/2) =
+             exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)"
+  using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex)
+
+end
+
+
+subsection \<open>Limits and residues\<close>
+
+text \<open>
+  The inverse of the Gamma function has simple zeros:
+\<close>
+
+lemma rGamma_zeros:
+  "(\<lambda>z. rGamma z / (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n * fact n :: 'a :: Gamma)"
+proof (subst tendsto_cong)
+  let ?f = "\<lambda>z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a"
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    show "eventually (\<lambda>z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))"
+    by (subst pochhammer_rGamma[of _ "Suc n"])
+       (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0)
+  have "isCont ?f (- of_nat n)" by (intro continuous_intros)
+  thus "?f \<midarrow> (- of_nat n) \<rightarrow> (- 1) ^ n * fact n" unfolding isCont_def
+    by (simp add: pochhammer_same)
+qed
+
+
+text \<open>
+  The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function,
+  and their residues can easily be computed from the limit we have just proven:
+\<close>
+
+lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))"
+proof -
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    have "eventually (\<lambda>z. rGamma z \<noteq> (0 :: 'a)) (at (- of_nat n))"
+    by (auto elim!: eventually_mono nonpos_Ints_cases'
+             simp: rGamma_eq_zero_iff dist_of_nat dist_minus)
+  with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident]
+    have "filterlim (\<lambda>z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))"
+    unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity])
+                            (simp_all add: filterlim_at)
+  moreover have "(\<lambda>z. inverse (rGamma z) :: 'a) = Gamma"
+    by (intro ext) (simp add: rGamma_inverse_Gamma)
+  ultimately show ?thesis by (simp only: )
+qed
+
+lemma Gamma_residues:
+  "(\<lambda>z. Gamma z * (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n / fact n :: 'a :: Gamma)"
+proof (subst tendsto_cong)
+  let ?c = "(- 1) ^ n / fact n :: 'a"
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    show "eventually (\<lambda>z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n)))
+            (at (- of_nat n))"
+    by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma)
+  have "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow>
+          inverse ((- 1) ^ n * fact n :: 'a)"
+    by (intro tendsto_intros rGamma_zeros) simp_all
+  also have "inverse ((- 1) ^ n * fact n) = ?c"
+    by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib)
+  finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" .
+qed
+
+
+
+subsection \<open>Alternative definitions\<close>
+
+
+subsubsection \<open>Variant of the Euler form\<close>
+
+
+definition Gamma_series_euler' where
+  "Gamma_series_euler' z n =
+     inverse z * (\<Prod>k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))"
+
+context
+begin
+private lemma Gamma_euler'_aux1:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes n: "n > 0"
+  shows "exp (z * of_real (ln (of_nat n + 1))) = (\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))"
+proof -
+  have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) =
+          exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
+    by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib)
+  also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
+    by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg)
+  also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
+    by (intro setprod.cong) (simp_all add: divide_simps)
+  also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1"
+    by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps)
+  finally show ?thesis ..
+qed
+
+lemma Gamma_series_euler':
+  assumes z: "(z :: 'a :: Gamma) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "(\<lambda>n. Gamma_series_euler' z n) \<longlonglongrightarrow> Gamma z"
+proof (rule Gamma_seriesI, rule Lim_transform_eventually)
+  let ?f = "\<lambda>n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)"
+  let ?r = "\<lambda>n. ?f n / Gamma_series z n"
+  let ?r' = "\<lambda>n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))"
+  from z have z': "z \<noteq> 0" by auto
+
+  have "eventually (\<lambda>n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"]
+    using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
+                     elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int)
+  moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))"
+    by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all
+  ultimately show "?r \<longlonglongrightarrow> 1" by (force dest!: Lim_transform_eventually)
+
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n z' have "Gamma_series_euler' z n =
+      exp (z * of_real (ln (of_nat n + 1))) / (z * (\<Prod>k=1..n. (1 + z / of_nat k)))"
+      by (subst Gamma_euler'_aux1)
+         (simp_all add: Gamma_series_euler'_def setprod.distrib
+                        setprod_inversef[symmetric] divide_inverse)
+    also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"
+      by (cases n) (simp_all add: pochhammer_setprod fact_setprod atLeastLessThanSuc_atLeastAtMost
+        setprod_dividef [symmetric] field_simps setprod.atLeast_Suc_atMost_Suc_shift)
+    also have "z * \<dots> = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec)
+    finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp
+  qed
+qed
+
+end
+
+
+
+subsubsection \<open>Weierstrass form\<close>
+
+definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series_weierstrass z n =
+     exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))"
+
+definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
+  "rGamma_series_weierstrass z n =
+     exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))"
+
+lemma Gamma_series_weierstrass_nonpos_Ints:
+  "eventually (\<lambda>k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially"
+  using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def)
+
+lemma rGamma_series_weierstrass_nonpos_Ints:
+  "eventually (\<lambda>k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially"
+  using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def)
+
+lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case True
+  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
+  also from True have "Gamma_series_weierstrass \<dots> \<longlonglongrightarrow> Gamma z"
+    by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int)
+  finally show ?thesis .
+next
+  case False
+  hence z: "z \<noteq> 0" by auto
+  let ?f = "(\<lambda>x. \<Prod>x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))"
+  have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \<ge> 1" for n :: nat
+    using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp)
+  have "(\<lambda>n. \<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \<longlonglongrightarrow> ln_Gamma z + euler_mascheroni * z + ln z"
+    using ln_Gamma_series'_aux[OF False]
+    by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def
+                   setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
+  from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"
+    by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A)
+  from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z
+    show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z"
+    by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def])
+qed
+
+lemma tendsto_complex_of_real_iff: "((\<lambda>x. complex_of_real (f x)) \<longlongrightarrow> of_real c) F = (f \<longlongrightarrow> c) F"
+  by (rule tendsto_of_real_iff)
+
+lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \<longlonglongrightarrow> Gamma (x :: real)"
+  using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def]
+  by (subst tendsto_complex_of_real_iff [symmetric])
+     (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real)
+
+lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \<longlonglongrightarrow> rGamma (z :: complex)"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case True
+  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
+  also from True have "rGamma_series_weierstrass \<dots> \<longlonglongrightarrow> rGamma z"
+    by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int)
+  finally show ?thesis .
+next
+  case False
+  have "rGamma_series_weierstrass z = (\<lambda>n. inverse (Gamma_series_weierstrass z n))"
+    by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def
+                  exp_minus divide_inverse setprod_inversef[symmetric] mult_ac)
+  also from False have "\<dots> \<longlonglongrightarrow> inverse (Gamma z)"
+    by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff)
+  finally show ?thesis by (simp add: Gamma_def)
+qed
+
+subsubsection \<open>Binomial coefficient form\<close>
+
+lemma Gamma_gbinomial:
+  "(\<lambda>n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \<longlonglongrightarrow> rGamma (z+1)"
+proof (cases "z = 0")
+  case False
+  show ?thesis
+  proof (rule Lim_transform_eventually)
+    let ?powr = "\<lambda>a b. exp (b * of_real (ln (of_nat a)))"
+    show "eventually (\<lambda>n. rGamma_series z n / z =
+            ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n"
+        by (simp add: gbinomial_pochhammer' pochhammer_rec)
+      also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z"
+        by (simp add: rGamma_series_def divide_simps exp_minus)
+      finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" ..
+    qed
+
+    from False have "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma z / z" by (intro tendsto_intros)
+    also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z]
+      by (simp add: field_simps)
+    finally show "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma (z+1)" .
+  qed
+qed (simp_all add: binomial_gbinomial [symmetric])
+
+lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)"
+  by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric])
+
+lemma gbinomial_asymptotic:
+  fixes z :: "'a :: Gamma"
+  shows "(\<lambda>n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \<longlonglongrightarrow>
+           inverse (Gamma (- z))"
+  unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"]
+  by (subst (asm) gbinomial_minus')
+     (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric])
+
+lemma fact_binomial_limit:
+  "(\<lambda>n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \<longlonglongrightarrow> 1 / fact k"
+proof (rule Lim_transform_eventually)
+  have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n))))
+            \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _")
+    using Gamma_gbinomial[of "of_nat k :: 'a"]
+    by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus)
+  also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact)
+  finally show "?f \<longlonglongrightarrow> 1 / fact k" .
+
+  show "eventually (\<lambda>n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)"
+      by (simp add: exp_of_nat_mult)
+    thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp
+  qed
+qed
+
+lemma binomial_asymptotic':
+  "(\<lambda>n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \<longlonglongrightarrow> 1"
+  using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp
+
+lemma gbinomial_Beta:
+  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))"
+using assms
+proof (induction n arbitrary: z)
+  case 0
+  hence "z + 2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute)
+  with 0 show ?case
+    by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute)
+next
+  case (Suc n z)
+  show ?case
+  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+    case True
+    with Suc.prems have "z = 0"
+      by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff)
+    show ?thesis
+    proof (cases "n = 0")
+      case True
+      with \<open>z = 0\<close> show ?thesis
+        by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric])
+    next
+      case False
+      with \<open>z = 0\<close> show ?thesis
+        by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1)
+    qed
+  next
+    case False
+    have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp
+    also have "\<dots> = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)"
+      by (subst gbinomial_factors) (simp add: field_simps)
+    also from False have "\<dots> = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))"
+      (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1)
+    also have "of_nat (Suc n) \<notin> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all
+    hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)"
+      by (subst Beta_plus1_right [symmetric]) simp_all
+    finally show ?thesis .
+  qed
+qed
+
+lemma gbinomial_Gamma:
+  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
+proof -
+  have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
+    by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac)
+  also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)"
+    using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac)
+  finally show ?thesis .
+qed
+
+
+subsubsection \<open>Integral form\<close>
+
+lemma integrable_Gamma_integral_bound:
+  fixes a c :: real
+  assumes a: "a > -1" and c: "c \<ge> 0"
+  defines "f \<equiv> \<lambda>x. if x \<in> {0..c} then x powr a else exp (-x/2)"
+  shows   "f integrable_on {0..}"
+proof -
+  have "f integrable_on {0..c}"
+    by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]])
+       (insert a c, simp_all add: f_def)
+  moreover have A: "(\<lambda>x. exp (-x/2)) integrable_on {c..}"
+    using integrable_on_exp_minus_to_infinity[of "1/2"] by simp
+  have "f integrable_on {c..}"
+    by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def)
+  ultimately show "f integrable_on {0..}"
+    by (rule integrable_union') (insert c, auto simp: max_def)
+qed
+
+lemma Gamma_integral_complex:
+  assumes z: "Re z > 0"
+  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
+proof -
+  have A: "((\<lambda>t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n))
+          has_integral (fact n / pochhammer z (n+1))) {0..1}"
+    if "Re z > 0" for n z using that
+  proof (induction n arbitrary: z)
+    case 0
+    have "((\<lambda>t. complex_of_real t powr (z - 1)) has_integral
+            (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0
+      by (intro fundamental_theorem_of_calculus_interior)
+         (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex)
+    thus ?case by simp
+  next
+    case (Suc n)
+    let ?f = "\<lambda>t. complex_of_real t powr z / z"
+    let ?f' = "\<lambda>t. complex_of_real t powr (z - 1)"
+    let ?g = "\<lambda>t. (1 - complex_of_real t) ^ Suc n"
+    let ?g' = "\<lambda>t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)"
+    have "((\<lambda>t. ?f' t * ?g t) has_integral
+            (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}"
+      (is "(_ has_integral ?I) _")
+    proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g])
+      from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g"
+        by (auto intro!: continuous_intros)
+    next
+      fix t :: real assume t: "t \<in> {0<..<1}"
+      show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems
+        by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex)
+      show "(?g has_vector_derivative ?g' t) (at t)"
+        by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all
+    next
+      from Suc.prems have [simp]: "z \<noteq> 0" by auto
+      from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp
+      have [simp]: "z + of_nat n \<noteq> 0" "z + 1 + of_nat n \<noteq> 0" for n
+        using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel)
+      have "((\<lambda>x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral
+              fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}"
+        (is "(?A has_integral ?B) _")
+        using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec)
+      also have "?A = (\<lambda>t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps)
+      also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))"
+        by (simp add: divide_simps pochhammer_rec
+              setprod_shift_bounds_cl_Suc_ivl del: of_nat_Suc)
+      finally show "((\<lambda>t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}"
+        by simp
+    qed (simp_all add: bounded_bilinear_mult)
+    thus ?case by simp
+  qed
+
+  have B: "((\<lambda>t. if t \<in> {0..of_nat n} then
+             of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0)
+           has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n
+  proof (cases "n > 0")
+    case [simp]: True
+    hence [simp]: "n \<noteq> 0" by auto
+    with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0]
+      have "((\<lambda>x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n)
+              has_integral fact n * of_nat n / pochhammer z (n+1)) ((\<lambda>x. real n * x)`{0..1})"
+      (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real)
+    also from True have "((\<lambda>x. real n*x)`{0..1}) = {0..real n}"
+      by (subst image_mult_atLeastAtMost) simp_all
+    also have "?f = (\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)"
+      using True by (intro ext) (simp add: field_simps)
+    finally have "((\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
+                    has_integral ?I) {0..real n}" (is ?P) .
+    also have "?P \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n)
+                        has_integral ?I) {0..real n}"
+      by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric])
+    also have "\<dots> \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n)
+                        has_integral ?I) {0..real n}"
+      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div)
+    finally have \<dots> .
+    note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"]
+    have "((\<lambda>x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n)
+            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P)
+      by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric])
+         (simp add: Ln_of_nat algebra_simps)
+    also have "?P \<longleftrightarrow> ((\<lambda>x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
+            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}"
+      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real)
+    also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) =
+                 (of_nat n powr z * fact n / pochhammer z (n+1))"
+      by (auto simp add: powr_def algebra_simps exp_diff)
+    finally show ?thesis by (subst has_integral_restrict) simp_all
+  next
+    case False
+    thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl)
+  qed
+
+  have "eventually (\<lambda>n. Gamma_series z n =
+          of_nat n powr z * fact n / pochhammer z (n+1)) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def)
+  from this and Gamma_series_LIMSEQ[of z]
+    have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z"
+    by (rule Lim_transform_eventually)
+
+  {
+    fix x :: real assume x: "x \<ge> 0"
+    have lim_exp: "(\<lambda>k. (1 - x / real k) ^ k) \<longlonglongrightarrow> exp (-x)"
+      using tendsto_exp_limit_sequentially[of "-x"] by simp
+    have "(\<lambda>k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k))
+            \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" (is ?P)
+      by (intro tendsto_intros lim_exp)
+    also from eventually_gt_at_top[of "nat \<lceil>x\<rceil>"]
+      have "eventually (\<lambda>k. of_nat k > x) sequentially" by eventually_elim linarith
+    hence "?P \<longleftrightarrow> (\<lambda>k. if x \<le> of_nat k then
+                 of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0)
+                   \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))"
+      by (intro tendsto_cong) (auto elim!: eventually_mono)
+    finally have \<dots> .
+  }
+  hence D: "\<forall>x\<in>{0..}. (\<lambda>k. if x \<in> {0..real k} then
+              of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0)
+             \<longlonglongrightarrow> of_real x powr (z - 1) / of_real (exp x)"
+    by (simp add: exp_minus field_simps cong: if_cong)
+
+  have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"
+    by (intro tendsto_intros ln_x_over_x_tendsto_0)
+  hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp
+  from order_tendstoD(2)[OF this, of "1/2"]
+    have "eventually (\<lambda>x. (Re z - 1) * ln x / x < 1/2) at_top" by simp
+  from eventually_conj[OF this eventually_gt_at_top[of 0]]
+    obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < 1/2 \<and> x > 0"
+    by (auto simp: eventually_at_top_linorder)
+  hence x0: "x0 > 0" "\<And>x. x \<ge> x0 \<Longrightarrow> (Re z - 1) * ln x < x / 2" by auto
+
+  define h where "h = (\<lambda>x. if x \<in> {0..x0} then x powr (Re z - 1) else exp (-x/2))"
+  have le_h: "x powr (Re z - 1) * exp (-x) \<le> h x" if x: "x \<ge> 0" for x
+  proof (cases "x > x0")
+    case True
+    from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)"
+      by (simp add: powr_def exp_diff exp_minus field_simps exp_add)
+    also from x0(2)[of x] True have "\<dots> < exp (-x/2)"
+      by (simp add: field_simps)
+    finally show ?thesis using True by (auto simp add: h_def)
+  next
+    case False
+    from x have "x powr (Re z - 1) * exp (- x) \<le> x powr (Re z - 1) * 1"
+      by (intro mult_left_mono) simp_all
+    with False show ?thesis by (auto simp add: h_def)
+  qed
+
+  have E: "\<forall>x\<in>{0..}. cmod (if x \<in> {0..real k} then of_real x powr (z - 1) *
+                   (1 - complex_of_real x / of_nat k) ^ k else 0) \<le> h x"
+    (is "\<forall>x\<in>_. ?f x \<le> _") for k
+  proof safe
+    fix x :: real assume x: "x \<ge> 0"
+    {
+      fix x :: real and n :: nat assume x: "x \<le> of_nat n"
+      have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp
+      also have "norm \<dots> = \<bar>(1 - x / real n)\<bar>" by (subst norm_of_real) (rule refl)
+      also from x have "\<dots> = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps)
+      finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" .
+    } note D = this
+    from D[of x k] x
+      have "?f x \<le> (if of_nat k \<ge> x \<and> k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)"
+      by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg)
+    also have "\<dots> \<le> x powr (Re z - 1) * exp  (-x)"
+      by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n)
+    also from x have "\<dots> \<le> h x" by (rule le_h)
+    finally show "?f x \<le> h x" .
+  qed
+
+  have F: "h integrable_on {0..}" unfolding h_def
+    by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all)
+  show ?thesis
+    by (rule has_integral_dominated_convergence[OF B F E D C])
+qed
+
+lemma Gamma_integral_real:
+  assumes x: "x > (0 :: real)"
+  shows   "((\<lambda>t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}"
+proof -
+  have A: "((\<lambda>t. complex_of_real t powr (complex_of_real x - 1) /
+          complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}"
+    using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real)
+  have "((\<lambda>t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}"
+    by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric])
+  from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def)
+qed
+
+
+
+subsection \<open>The Weierstraß product formula for the sine\<close>
+
+lemma sin_product_formula_complex:
+  fixes z :: complex
+  shows "(\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k^2)) \<longlonglongrightarrow> sin (of_real pi * z)"
+proof -
+  let ?f = "rGamma_series_weierstrass"
+  have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n))
+            \<longlonglongrightarrow> (- of_real pi * inverse z) * (rGamma z * rGamma (- z))"
+    by (intro tendsto_intros rGamma_weierstrass_complex)
+  also have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) =
+                    (\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2))"
+  proof
+    fix n :: nat
+    have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) =
+              of_real pi * z * (\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)"
+      by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus
+                    divide_simps setprod.distrib[symmetric] power2_eq_square)
+    also have "(\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) =
+                 (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)"
+      by (intro setprod.cong) (simp_all add: power2_eq_square field_simps)
+    finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \<dots>"
+      by (simp add: divide_simps)
+  qed
+  also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)"
+    by (subst rGamma_reflection_complex') (simp add: divide_simps)
+  finally show ?thesis .
+qed
+
+lemma sin_product_formula_real:
+  "(\<lambda>n. pi * (x::real) * (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x)"
+proof -
+  from sin_product_formula_complex[of "of_real x"]
+    have "(\<lambda>n. of_real pi * of_real x * (\<Prod>k=1..n. 1 - (of_real x)^2 / (of_nat k)^2))
+              \<longlonglongrightarrow> sin (of_real pi * of_real x :: complex)" (is "?f \<longlonglongrightarrow> ?y") .
+  also have "?f = (\<lambda>n. of_real (pi * x * (\<Prod>k=1..n. 1 - x^2 / (of_nat k^2))))" by simp
+  also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult)
+  finally show ?thesis by (subst (asm) tendsto_of_real_iff)
+qed
+
+lemma sin_product_formula_real':
+  assumes "x \<noteq> (0::real)"
+  shows   "(\<lambda>n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x) / (pi * x)"
+  using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms
+  by simp
+
+
+subsection \<open>The Solution to the Basel problem\<close>
+
+theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)"
+proof -
+  define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n
+  define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)"
+  define f where [abs_def]: "f x = (\<Sum>n. P x n / of_nat (Suc n)^2)" for x
+  define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x
+
+  have sums: "(\<lambda>n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x
+  proof (cases "x = 0")
+    assume x: "x = 0"
+    have "summable (\<lambda>n. inverse ((real_of_nat (Suc n))\<^sup>2))"
+      using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp
+    thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums)
+  next
+    assume x: "x \<noteq> 0"
+    have "(\<lambda>n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))"
+      unfolding P_def using x by (intro telescope_sums' sin_product_formula_real')
+    also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)"
+      unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps)
+    also have "P x 0 = 1" by (simp add: P_def)
+    finally have "(\<lambda>n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" .
+    from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp
+  qed
+
+  have "continuous_on (ball 0 1) f"
+  proof (rule uniform_limit_theorem; (intro always_eventually allI)?)
+    show "uniform_limit (ball 0 1) (\<lambda>n x. \<Sum>k<n. P x k / of_nat (Suc k)^2) f sequentially"
+    proof (unfold f_def, rule weierstrass_m_test)
+      fix n :: nat and x :: real assume x: "x \<in> ball 0 1"
+      {
+        fix k :: nat assume k: "k \<ge> 1"
+        from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1)
+        also from k have "\<dots> \<le> of_nat k^2" by simp
+        finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k
+          by (simp_all add: field_simps del: of_nat_Suc)
+      }
+      hence "(\<Prod>k=1..n. abs (1 - x^2 / of_nat k^2)) \<le> (\<Prod>k=1..n. 1)" by (intro setprod_mono) simp
+      thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2"
+        unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc)
+    qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide)
+  qed (auto simp: P_def intro!: continuous_intros)
+  hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all
+  hence "(f \<midarrow> 0 \<rightarrow> f 0)" by (simp add: isCont_def)
+  also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide)
+  finally have "f \<midarrow> 0 \<rightarrow> K" .
+
+  moreover have "f \<midarrow> 0 \<rightarrow> pi^2 / 6"
+  proof (rule Lim_transform_eventually)
+    define f' where [abs_def]: "f' x = (\<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x
+    have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)"
+      by (auto simp add: eventually_at intro!: exI[of _ 1])
+    thus "eventually (\<lambda>x. f' x = f x) (at 0)"
+    proof eventually_elim
+      fix x :: real assume x: "x \<noteq> 0"
+      have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def)
+      with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"]
+      have "(\<lambda>n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))"
+        by (simp add: eval_nat_numeral)
+      from sums_divide[OF this, of "x^3 * pi"] x
+        have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)"
+        by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac)
+      with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)"
+        by (simp add: g_def)
+      hence "f' x = g x / x^2" by (simp add: sums_iff f'_def)
+      also have "\<dots> = f x" using sums[of x] x by (simp add: sums_iff g_def f_def)
+      finally show "f' x = f x" .
+    qed
+
+    have "isCont f' 0" unfolding f'_def
+    proof (intro isCont_powser_converges_everywhere)
+      fix x :: real show "summable (\<lambda>n. -sin_coeff (n+3) * pi^(n+2) * x^n)"
+      proof (cases "x = 0")
+        assume x: "x \<noteq> 0"
+        from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF
+               sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x
+          show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral)
+      qed (simp only: summable_0_powser)
+    qed
+    hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def)
+    also have "f' 0 = pi * pi / fact 3" unfolding f'_def
+      by (subst powser_zero) (simp add: sin_coeff_def)
+    finally show "f' \<midarrow> 0 \<rightarrow> pi^2 / 6" by (simp add: eval_nat_numeral)
+  qed
+
+  ultimately have "K = pi^2 / 6" by (rule LIM_unique)
+  moreover from inverse_power_summable[of 2]
+    have "summable (\<lambda>n. (inverse (real_of_nat (Suc n)))\<^sup>2)"
+    by (subst summable_Suc_iff) (simp add: power_inverse)
+  ultimately show ?thesis unfolding K_def
+    by (auto simp add: sums_iff power_divide inverse_eq_divide)
+qed
+
+end
--- a/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -11,10 +11,10 @@
 \<close>
 
 theory Generalised_Binomial_Theorem
-imports 
-  Complex_Main 
+imports
+  Complex_Main
   Complex_Transcendental
-  Summation
+  Summation_Tests
 begin
 
 lemma gbinomial_ratio_limit:
@@ -36,14 +36,14 @@
       by (simp add: setprod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
     also have "?P / \<dots> = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric])
     also from assms have "?P / ?P = 1" by auto
-    also have "of_nat (Suc n) * (1 / (a - of_nat n)) = 
+    also have "of_nat (Suc n) * (1 / (a - of_nat n)) =
                    inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps)
     also have "inverse (of_nat (Suc n)) * (a - of_nat n) = a / of_nat (Suc n) - of_nat n / of_nat (Suc n)"
       by (simp add: field_simps del: of_nat_Suc)
     finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp
   qed
 
-  have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0" 
+  have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0"
     unfolding divide_inverse
     by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat)
   hence "(\<lambda>n. a / of_nat (Suc n)) \<longlonglongrightarrow> 0"
@@ -84,26 +84,26 @@
   with K have summable: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < K" for z using that by auto
   hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that
     by (intro termdiff_converges[of _ K]) simp_all
-  
+
   define f f' where [abs_def]: "f z = (\<Sum>n. ?f n * z ^ n)" "f' z = (\<Sum>n. ?f' n * z ^ n)" for z
   {
     fix z :: complex assume z: "norm z < K"
     from summable_mult2[OF summable'[OF z], of z]
       have summable1: "summable (\<lambda>n. ?f' n * z ^ Suc n)" by (simp add: mult_ac)
-    hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)" 
+    hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)"
       unfolding diffs_def by (subst (asm) summable_Suc_iff)
 
     have "(1 + z) * f' z = (\<Sum>n. ?f' n * z^n) + (\<Sum>n. ?f' n * z^Suc n)"
       unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult)
     also have "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)"
       by (intro suminf_cong) (simp add: diffs_def)
-    also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)" 
+    also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)"
       using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def
       by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all
     also have "(\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\<Sum>n. of_nat n * ?f n * z^n) =
                  (\<Sum>n. a * ?f n * z^n)"
       by (subst gbinomial_mult_1, subst suminf_add)
-         (insert summable'[OF z] summable2, 
+         (insert summable'[OF z] summable2,
           simp_all add: summable_powser_split_head algebra_simps diffs_def)
     also have "\<dots> = a * f z" unfolding f_f'_def
       by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac)
@@ -124,16 +124,16 @@
     with K have nz: "1 + z \<noteq> 0" by (auto dest!: minus_unique)
     from z K have "norm z < 1" by simp
     hence "(1 + z) \<notin> \<real>\<^sub>\<le>\<^sub>0" by (cases z) (auto simp: complex_nonpos_Reals_iff)
-    hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 
+    hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative
               f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z
       by (auto intro!: derivative_eq_intros)
     also from z have "a * f z = (1 + z) * f' z" by (rule deriv)
-    finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)" 
+    finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)"
       using nz by (simp add: field_simps powr_diff_complex at_within_open[OF z'])
   qed simp_all
   then obtain c where c: "\<And>z. z \<in> ball 0 K \<Longrightarrow> f z * (1 + z) powr (-a) = c" by blast
   from c[of 0] and K have "c = 1" by simp
-  with c[of z] have "f z = (1 + z) powr a" using K 
+  with c[of z] have "f z = (1 + z) powr a" using K
     by (simp add: powr_minus_complex field_simps dist_complex_def)
   with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff)
 qed
@@ -141,7 +141,7 @@
 lemma gen_binomial_complex':
   fixes x y :: real and a :: complex
   assumes "\<bar>x\<bar> < \<bar>y\<bar>"
-  shows   "(\<lambda>n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums 
+  shows   "(\<lambda>n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums
                of_real (x + y) powr a" (is "?P x y")
 proof -
   {
@@ -150,7 +150,7 @@
     note xy = xy this
     from xy have "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n) sums (1 + of_real (x / y)) powr a"
         by (intro gen_binomial_complex) (simp add: norm_divide)
-    hence "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums 
+    hence "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums
                ((1 + of_real (x / y)) powr a * y powr a)"
       by (rule sums_mult2)
     also have "(1 + complex_of_real (x / y)) = complex_of_real (1 + x/y)" by simp
@@ -172,7 +172,7 @@
       by (subst powr_neg_real_complex) (simp add: abs_real_def split: if_split_asm)
     also {
       fix n :: nat
-      from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) = 
+      from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) =
                        (a gchoose n) * (-of_real x / -of_real y) ^ n * (- of_real y) powr a"
         by (subst power_divide) (simp add: powr_diff_complex powr_nat)
       also from y have "(- of_real y) powr a = (-1) powr -a * of_real y powr a"
@@ -180,7 +180,7 @@
       also have "-complex_of_real x / -complex_of_real y = complex_of_real x / complex_of_real y"
         by simp
       also have "... ^ n = of_real x ^ n / of_real y ^ n" by (simp add: power_divide)
-      also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) = 
+      also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) =
                    (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - n))"
         by (simp add: algebra_simps powr_diff_complex powr_nat)
       finally have "(a gchoose n) * of_real (- x) ^ n * of_real (- y) powr (a - of_nat n) =
@@ -194,7 +194,7 @@
 lemma gen_binomial_complex'':
   fixes x y :: real and a :: complex
   assumes "\<bar>y\<bar> < \<bar>x\<bar>"
-  shows   "(\<lambda>n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums 
+  shows   "(\<lambda>n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums
                of_real (x + y) powr a"
   using gen_binomial_complex'[OF assms] by (simp add: mult_ac add.commute)
 
@@ -209,12 +209,12 @@
               (of_real (1 + z)) powr (of_real a)" by simp
   also have "(of_real (1 + z) :: complex) powr (of_real a) = of_real ((1 + z) powr a)"
     using assms by (subst powr_of_real) simp_all
-  also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n 
+  also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n
     by (simp add: gbinomial_setprod_rev)
   hence "(\<lambda>n. (of_real a gchoose n :: complex) * of_real z ^ n) =
            (\<lambda>n. of_real ((a gchoose n) * z ^ n))" by (intro ext) simp
   finally show ?thesis by (simp only: sums_of_real_iff)
-qed 
+qed
 
 lemma gen_binomial_real':
   fixes x y a :: real
@@ -228,7 +228,7 @@
     by (rule gen_binomial_real)
   hence "(\<lambda>n. (a gchoose n) * (x / y) ^ n * y powr a) sums ((1 + x / y) powr a * y powr a)"
     by (rule sums_mult2)
-  with xy show ?thesis 
+  with xy show ?thesis
     by (simp add: field_simps powr_divide powr_divide2[symmetric] powr_realpow)
 qed
 
@@ -245,7 +245,7 @@
   using gen_binomial_real'[OF assms] by (simp add: mult_ac add.commute)
 
 lemma sqrt_series':
-  "\<bar>z\<bar> < a \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums 
+  "\<bar>z\<bar> < a \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums
                   sqrt (a + z :: real)"
   using gen_binomial_real''[of z a "1/2"] by (simp add: powr_half_sqrt)
 
--- a/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -5,15 +5,15 @@
 section \<open>Harmonic Numbers\<close>
 
 theory Harmonic_Numbers
-imports 
+imports
   Complex_Transcendental
-  Summation
+  Summation_Tests
   Integral_Test
 begin
 
 text \<open>
   The definition of the Harmonic Numbers and the Euler-Mascheroni constant.
-  Also provides a reasonably accurate approximation of @{term "ln 2 :: real"} 
+  Also provides a reasonably accurate approximation of @{term "ln 2 :: real"}
   and the Euler-Mascheroni constant.
 \<close>
 
@@ -51,11 +51,11 @@
 
 lemma of_real_harm: "of_real (harm n) = harm n"
   unfolding harm_def by simp
-  
+
 lemma norm_harm: "norm (harm n) = harm n"
   by (subst of_real_harm [symmetric]) (simp add: harm_nonneg)
 
-lemma harm_expand: 
+lemma harm_expand:
   "harm 0 = 0"
   "harm (Suc 0) = 1"
   "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)"
@@ -134,7 +134,7 @@
            has_field_derivative_iff_has_vector_derivative[symmetric])
   hence "integral {0..of_nat n} (\<lambda>x. inverse (x + 1) :: real) = ln (of_nat (Suc n))"
     by (auto dest!: integral_unique)
-  ultimately show ?thesis 
+  ultimately show ?thesis
     by (simp add: euler_mascheroni.sum_integral_diff_series_def atLeast0AtMost)
 qed
 
@@ -151,7 +151,7 @@
 
 lemma euler_mascheroni_convergent: "convergent (\<lambda>n. harm n - ln (of_nat n) :: real)"
 proof -
-  have A: "(\<lambda>n. harm (Suc n) - ln (of_nat (Suc n))) = 
+  have A: "(\<lambda>n. harm (Suc n) - ln (of_nat (Suc n))) =
              euler_mascheroni.sum_integral_diff_series"
     by (subst euler_mascheroni_sum_integral_diff_series [symmetric]) (rule refl)
   have "convergent (\<lambda>n. harm (Suc n) - ln (of_nat (Suc n) :: real))"
@@ -159,13 +159,13 @@
   thus ?thesis by (subst (asm) convergent_Suc_iff)
 qed
 
-lemma euler_mascheroni_LIMSEQ: 
+lemma euler_mascheroni_LIMSEQ:
   "(\<lambda>n. harm n - ln (of_nat n) :: real) \<longlonglongrightarrow> euler_mascheroni"
   unfolding euler_mascheroni_def
   by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent)
 
-lemma euler_mascheroni_LIMSEQ_of_real: 
-  "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> 
+lemma euler_mascheroni_LIMSEQ_of_real:
+  "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow>
       (euler_mascheroni :: 'a :: {real_normed_algebra_1, topological_space})"
 proof -
   have "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> (of_real (euler_mascheroni) :: 'a)"
@@ -175,7 +175,7 @@
 
 lemma euler_mascheroni_sum:
   "(\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real)
-       sums euler_mascheroni" 
+       sums euler_mascheroni"
  using sums_add[OF telescope_sums[OF LIMSEQ_Suc[OF euler_mascheroni_LIMSEQ]]
                    telescope_sums'[OF LIMSEQ_inverse_real_of_nat]]
   by (simp_all add: harm_def algebra_simps)
@@ -198,21 +198,21 @@
       by (intro setsum.mono_neutral_right) auto
     also have "\<dots> = (\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k)))"
       by (intro setsum.cong) auto
-    also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n" 
+    also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n"
       unfolding harm_altdef
       by (intro setsum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE)
     also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n
       by (simp_all add: algebra_simps ln_mult)
     finally show "?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))" ..
   qed
-  moreover have "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) 
+  moreover have "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real))
                      \<longlonglongrightarrow> euler_mascheroni - euler_mascheroni + ln 2"
     by (intro tendsto_intros euler_mascheroni_LIMSEQ filterlim_compose[OF euler_mascheroni_LIMSEQ]
               filterlim_subseq) (auto simp: subseq_def)
   hence "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) \<longlonglongrightarrow> ln 2" by simp
   ultimately have "(\<lambda>n. (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
     by (rule Lim_transform_eventually)
-  
+
   moreover have "summable (\<lambda>k. (-1)^k * inverse (real_of_nat (Suc k)))"
     using LIMSEQ_inverse_real_of_nat
     by (intro summable_Leibniz(1) decseq_imp_monoseq decseq_SucI) simp_all
@@ -225,12 +225,12 @@
   with A show ?thesis by (simp add: sums_def)
 qed
 
-lemma alternating_harmonic_series_sums': 
+lemma alternating_harmonic_series_sums':
   "(\<lambda>k. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2))) sums ln 2"
 unfolding sums_def
 proof (rule Lim_transform_eventually)
   show "(\<lambda>n. \<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
-    using alternating_harmonic_series_sums unfolding sums_def 
+    using alternating_harmonic_series_sums unfolding sums_def
     by (rule filterlim_compose) (rule mult_nat_left_at_top, simp)
   show "eventually (\<lambda>n. (\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) =
             (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))) sequentially"
@@ -240,7 +240,7 @@
               (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))"
       by (induction n) (simp_all add: inverse_eq_divide)
   qed
-qed               
+qed
 
 
 subsection \<open>Bounds on the Euler--Mascheroni constant\<close>
@@ -254,16 +254,16 @@
   have f'_nonpos: "f' \<le> 0" using assms by (simp add: f'_def divide_simps)
   let ?f = "\<lambda>t. (t - x) * f' + inverse x"
   let ?F = "\<lambda>t. (t - x)^2 * f' / 2 + t * inverse x"
-  have diff: "\<forall>t\<in>{x..x+a}. (?F has_vector_derivative ?f t) 
+  have diff: "\<forall>t\<in>{x..x+a}. (?F has_vector_derivative ?f t)
                                (at t within {x..x+a})" using assms
-    by (auto intro!: derivative_eq_intros 
+    by (auto intro!: derivative_eq_intros
              simp: has_field_derivative_iff_has_vector_derivative[symmetric])
   from assms have "(?f has_integral (?F (x+a) - ?F x)) {x..x+a}"
     by (intro fundamental_theorem_of_calculus[OF _ diff])
        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] field_simps
              intro!: derivative_eq_intros)
   also have "?F (x+a) - ?F x = (a*2 + f'*a\<^sup>2*x) / (2*x)" using assms by (simp add: field_simps)
-  also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms 
+  also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms
     by (simp add: divide_simps f'_def power2_eq_square)
   also have "(a*2 + - a\<^sup>2/(x*(x+a))*x) / (2*x) = ?A" using assms
     by (simp add: divide_simps power2_eq_square) (simp add: algebra_simps)
@@ -281,7 +281,7 @@
     have "inverse t = inverse ((1 - (t - x) / a) *\<^sub>R x + ((t - x) / a) *\<^sub>R (x + a))" (is "_ = ?A")
       using assms t' by (simp add: field_simps)
     also from assms have "convex_on {x..x+a} inverse" by (intro convex_on_inverse) auto
-    from convex_onD_Icc[OF this _ t] assms 
+    from convex_onD_Icc[OF this _ t] assms
       have "?A \<le> (1 - (t - x) / a) * inverse x + (t - x) / a * inverse (x + a)" by simp
     also have "\<dots> = (t - x) * f' + inverse x" using assms
       by (simp add: f'_def divide_simps) (simp add: f'_def field_simps)
@@ -305,7 +305,7 @@
     by (intro fundamental_theorem_of_calculus)
        (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
              intro!: derivative_eq_intros)
-  also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m" 
+  also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m"
     by (simp add: field_simps)
   also have "((y - m)^2 - (x - m)^2) = 0" by (simp add: m_def power2_eq_square field_simps)
   also have "0 * f' / 2 + (y - x) / m = ?A" by (simp add: m_def)
@@ -332,7 +332,7 @@
 qed
 
 
-lemma euler_mascheroni_lower: 
+lemma euler_mascheroni_lower:
         "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
   and euler_mascheroni_upper:
         "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
@@ -370,11 +370,11 @@
     also have "\<dots> = D k" unfolding D_def inv_def ..
     finally show "D (k' + Suc n) \<ge> (inv (k' + Suc n + 1) - inv (k' + Suc n + 2)) / 2"
       by (simp add: k_def)
-    from sums_summable[OF sums] 
+    from sums_summable[OF sums]
       show "summable (\<lambda>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2))/2)" by simp
   qed
   also from sums have "\<dots> = -inv (n+2) / 2" by (simp add: sums_iff)
-  finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))" 
+  finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))"
     by (simp add: inv_def field_simps)
   also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
     unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
@@ -382,7 +382,7 @@
     by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
   finally show "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
     by simp
-  
+
   note sum
   also have "-(\<Sum>k. D (k + Suc n)) \<ge> -(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)"
   proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable])
@@ -393,22 +393,22 @@
     from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"]
       have "2 / (2 * real_of_nat k + 3) \<le> ln (of_nat (k+2)) - ln (real_of_nat (k+1))"
       by (simp add: add_ac)
-    hence "D k \<le> 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)" 
+    hence "D k \<le> 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)"
       by (simp add: D_def inverse_eq_divide inv_def)
     also have "\<dots> = inv ((k+1)*(2*k+3))" unfolding inv_def by (simp add: field_simps)
     also have "\<dots> \<le> inv (2*k*(k+1))" unfolding inv_def using k
-      by (intro le_imp_inverse_le) 
+      by (intro le_imp_inverse_le)
          (simp add: algebra_simps, simp del: of_nat_add)
     also have "\<dots> = (inv k - inv (k+1))/2" unfolding inv_def using k
       by (simp add: divide_simps del: of_nat_mult) (simp add: algebra_simps)
     finally show "D k \<le> (inv (Suc (k' + n)) - inv (Suc (Suc k' + n)))/2" unfolding k_def by simp
   next
-    from sums_summable[OF sums'] 
+    from sums_summable[OF sums']
       show "summable (\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" by simp
   qed
   also from sums' have "(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) = inv (n+1)/2"
     by (simp add: sums_iff)
-  finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))" 
+  finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))"
     by (simp add: inv_def field_simps)
   also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
     unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
@@ -428,7 +428,7 @@
   fixes n :: nat and x :: real
   defines "y \<equiv> (x-1)/(x+1)"
   assumes x: "x > 0" "x \<noteq> 1"
-  shows "inverse (2*y^(2*n+1)) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in> 
+  shows "inverse (2*y^(2*n+1)) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
             {0..(1 / (1 - y^2) / of_nat (2*n+1))}"
 proof -
   from x have norm_y: "norm y < 1" unfolding y_def by simp
@@ -468,21 +468,21 @@
   and     ln_approx_abs:    "abs (ln x - (approx + d)) \<le> d"
 proof -
   define c where "c = 2*y^(2*n+1)"
-  from x have c_pos: "c > 0" unfolding c_def y_def 
+  from x have c_pos: "c > 0" unfolding c_def y_def
     by (intro mult_pos_pos zero_less_power) simp_all
   have A: "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
               {0.. (1 / (1 - y^2) / of_nat (2*n+1))}" using assms unfolding y_def c_def
     by (intro ln_approx_aux) simp_all
   hence "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1)/of_nat (2*k+1))) \<le> (1 / (1-y^2) / of_nat (2*n+1))"
     by simp
-  hence "(ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) / c \<le> (1 / (1 - y^2) / of_nat (2*n+1))" 
+  hence "(ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) / c \<le> (1 / (1 - y^2) / of_nat (2*n+1))"
     by (auto simp add: divide_simps)
   with c_pos have "ln x \<le> c / (1 - y^2) / of_nat (2*n+1) + approx"
     by (subst (asm) pos_divide_le_eq) (simp_all add: mult_ac approx_def)
   moreover {
     from A c_pos have "0 \<le> c * (inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))))"
       by (intro mult_nonneg_nonneg[of c]) simp_all
-    also have "\<dots> = (c * inverse c) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1)))"  
+    also have "\<dots> = (c * inverse c) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1)))"
       by (simp add: mult_ac)
     also from c_pos have "c * inverse c = 1" by simp
     finally have "ln x \<ge> approx" by (simp add: approx_def)
@@ -501,7 +501,7 @@
 
 lemma euler_mascheroni_bounds':
   fixes n :: nat assumes "n \<ge> 1" "ln (real_of_nat (Suc n)) \<in> {l<..<u}"
-  shows "euler_mascheroni \<in> 
+  shows "euler_mascheroni \<in>
            {harm n - u + inverse (of_nat (2*(n+1)))<..<harm n - l + inverse (of_nat (2*n))}"
   using euler_mascheroni_bounds[OF assms(1)] assms(2) by auto
 
@@ -510,13 +510,13 @@
   Approximation of @{term "ln 2"}. The lower bound is accurate to about 0.03; the upper
   bound is accurate to about 0.0015.
 \<close>
-lemma ln2_ge_two_thirds: "2/3 \<le> ln (2::real)" 
+lemma ln2_ge_two_thirds: "2/3 \<le> ln (2::real)"
   and ln2_le_25_over_36: "ln (2::real) \<le> 25/36"
   using ln_approx_bounds[of 2 1, simplified, simplified eval_nat_numeral, simplified] by simp_all
 
 
 text \<open>
-  Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015; 
+  Approximation of the Euler--Mascheroni constant. The lower bound is accurate to about 0.0015;
   the upper bound is accurate to about 0.015.
 \<close>
 lemma euler_mascheroni_gt_19_over_33: "(euler_mascheroni :: real) > 19/33" (is ?th1)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Henstock_Kurzweil_Integration.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,12645 @@
+(*  Author:     John Harrison
+    Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP
+*)
+
+section \<open>Henstock-Kurzweil gauge integration in many dimensions.\<close>
+
+theory Henstock_Kurzweil_Integration
+imports
+  Derivative
+  Uniform_Limit
+  "~~/src/HOL/Library/Indicator_Function"
+begin
+
+lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
+  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
+  scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
+
+
+subsection \<open>Sundries\<close>
+
+lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
+lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
+lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto
+
+declare norm_triangle_ineq4[intro]
+
+lemma simple_image: "{f x |x . x \<in> s} = f ` s"
+  by blast
+
+lemma linear_simps:
+  assumes "bounded_linear f"
+  shows
+    "f (a + b) = f a + f b"
+    "f (a - b) = f a - f b"
+    "f 0 = 0"
+    "f (- a) = - f a"
+    "f (s *\<^sub>R v) = s *\<^sub>R (f v)"
+proof -
+  interpret f: bounded_linear f by fact
+  show "f (a + b) = f a + f b" by (rule f.add)
+  show "f (a - b) = f a - f b" by (rule f.diff)
+  show "f 0 = 0" by (rule f.zero)
+  show "f (- a) = - f a" by (rule f.minus)
+  show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR)
+qed
+
+lemma bounded_linearI:
+  assumes "\<And>x y. f (x + y) = f x + f y"
+    and "\<And>r x. f (r *\<^sub>R x) = r *\<^sub>R f x"
+    and "\<And>x. norm (f x) \<le> norm x * K"
+  shows "bounded_linear f"
+  using assms by (rule bounded_linear_intro) (* FIXME: duplicate *)
+
+lemma bounded_linear_component [intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x \<bullet> k)"
+  by (rule bounded_linear_inner_left)
+
+lemma transitive_stepwise_lt_eq:
+  assumes "(\<And>x y z::nat. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z)"
+  shows "((\<forall>m. \<forall>n>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n)))"
+  (is "?l = ?r")
+proof safe
+  assume ?r
+  fix n m :: nat
+  assume "m < n"
+  then show "R m n"
+  proof (induct n arbitrary: m)
+    case 0
+    then show ?case by auto
+  next
+    case (Suc n)
+    show ?case
+    proof (cases "m < n")
+      case True
+      show ?thesis
+        apply (rule assms[OF Suc(1)[OF True]])
+        using \<open>?r\<close>
+        apply auto
+        done
+    next
+      case False
+      then have "m = n"
+        using Suc(2) by auto
+      then show ?thesis
+        using \<open>?r\<close> by auto
+    qed
+  qed
+qed auto
+
+lemma transitive_stepwise_gt:
+  assumes "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" "\<And>n. R n (Suc n)"
+  shows "\<forall>n>m. R m n"
+proof -
+  have "\<forall>m. \<forall>n>m. R m n"
+    apply (subst transitive_stepwise_lt_eq)
+    apply (blast intro: assms)+
+    done
+  then show ?thesis by auto
+qed
+
+lemma transitive_stepwise_le_eq:
+  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
+  shows "(\<forall>m. \<forall>n\<ge>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n))"
+  (is "?l = ?r")
+proof safe
+  assume ?r
+  fix m n :: nat
+  assume "m \<le> n"
+  then show "R m n"
+  proof (induct n arbitrary: m)
+    case 0
+    with assms show ?case by auto
+  next
+    case (Suc n)
+    show ?case
+    proof (cases "m \<le> n")
+      case True
+      with Suc.hyps \<open>\<forall>n. R n (Suc n)\<close> assms show ?thesis
+        by blast
+    next
+      case False
+      then have "m = Suc n"
+        using Suc(2) by auto
+      then show ?thesis
+        using assms(1) by auto
+    qed
+  qed
+qed auto
+
+lemma transitive_stepwise_le:
+  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
+    and "\<And>n. R n (Suc n)"
+  shows "\<forall>n\<ge>m. R m n"
+proof -
+  have "\<forall>m. \<forall>n\<ge>m. R m n"
+    apply (subst transitive_stepwise_le_eq)
+    apply (blast intro: assms)+
+    done
+  then show ?thesis by auto
+qed
+
+
+subsection \<open>Some useful lemmas about intervals.\<close>
+
+lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
+  using nonempty_Basis
+  by (fastforce simp add: set_eq_iff mem_box)
+
+lemma interior_subset_union_intervals:
+  assumes "i = cbox a b"
+    and "j = cbox c d"
+    and "interior j \<noteq> {}"
+    and "i \<subseteq> j \<union> s"
+    and "interior i \<inter> interior j = {}"
+  shows "interior i \<subseteq> interior s"
+proof -
+  have "box a b \<inter> cbox c d = {}"
+     using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
+     unfolding assms(1,2) interior_cbox by auto
+  moreover
+  have "box a b \<subseteq> cbox c d \<union> s"
+    apply (rule order_trans,rule box_subset_cbox)
+    using assms(4) unfolding assms(1,2)
+    apply auto
+    done
+  ultimately
+  show ?thesis
+    unfolding assms interior_cbox
+      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
+qed
+
+lemma inter_interior_unions_intervals:
+  fixes f::"('a::euclidean_space) set set"
+  assumes "finite f"
+    and "open s"
+    and "\<forall>t\<in>f. \<exists>a b. t = cbox a b"
+    and "\<forall>t\<in>f. s \<inter> (interior t) = {}"
+  shows "s \<inter> interior (\<Union>f) = {}"
+proof (clarsimp simp only: all_not_in_conv [symmetric])
+  fix x
+  assume x: "x \<in> s" "x \<in> interior (\<Union>f)"
+  have lem1: "\<And>x e s U. ball x e \<subseteq> s \<inter> interior U \<longleftrightarrow> ball x e \<subseteq> s \<inter> U"
+    using interior_subset
+    by auto (meson Topology_Euclidean_Space.open_ball contra_subsetD interior_maximal mem_ball)
+  have "\<exists>t\<in>f. \<exists>x. \<exists>e>0. ball x e \<subseteq> s \<inter> t"
+    if "finite f" and "\<forall>t\<in>f. \<exists>a b. t = cbox a b" and "\<exists>x. x \<in> s \<inter> interior (\<Union>f)" for f
+    using that
+  proof (induct rule: finite_induct)
+    case empty
+    obtain x where "x \<in> s \<inter> interior (\<Union>{})"
+      using empty(2) ..
+    then have False
+      unfolding Union_empty interior_empty by auto
+    then show ?case by auto
+  next
+    case (insert i f)
+    obtain x where x: "x \<in> s \<inter> interior (\<Union>insert i f)"
+      using insert(5) ..
+    then obtain e where e: "0 < e \<and> ball x e \<subseteq> s \<inter> interior (\<Union>insert i f)"
+      unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior], rule_format] ..
+    obtain a where "\<exists>b. i = cbox a b"
+      using insert(4)[rule_format,OF insertI1] ..
+    then obtain b where ab: "i = cbox a b" ..
+    show ?case
+    proof (cases "x \<in> i")
+      case False
+      then have "x \<in> UNIV - cbox a b"
+        unfolding ab by auto
+      then obtain d where "0 < d \<and> ball x d \<subseteq> UNIV - cbox a b"
+        unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_cbox],rule_format] ..
+      then have "0 < d" "ball x (min d e) \<subseteq> UNIV - i"
+        unfolding ab ball_min_Int by auto
+      then have "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)"
+        using e unfolding lem1 unfolding  ball_min_Int by auto
+      then have "x \<in> s \<inter> interior (\<Union>f)" using \<open>d>0\<close> e by auto
+      then have "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t"
+        using insert.hyps(3) insert.prems(1) by blast
+      then show ?thesis by auto
+    next
+      case True show ?thesis
+      proof (cases "x\<in>box a b")
+        case True
+        then obtain d where "0 < d \<and> ball x d \<subseteq> box a b"
+          unfolding open_contains_ball_eq[OF open_box,rule_format] ..
+        then show ?thesis
+          apply (rule_tac x=i in bexI, rule_tac x=x in exI, rule_tac x="min d e" in exI)
+          unfolding ab
+          using box_subset_cbox[of a b] and e
+          apply fastforce+
+          done
+      next
+        case False
+        then obtain k where "x\<bullet>k \<le> a\<bullet>k \<or> x\<bullet>k \<ge> b\<bullet>k" and k: "k \<in> Basis"
+          unfolding mem_box by (auto simp add: not_less)
+        then have "x\<bullet>k = a\<bullet>k \<or> x\<bullet>k = b\<bullet>k"
+          using True unfolding ab and mem_box
+            apply (erule_tac x = k in ballE)
+            apply auto
+            done
+        then have "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)"
+        proof (rule disjE)
+          let ?z = "x - (e/2) *\<^sub>R k"
+          assume as: "x\<bullet>k = a\<bullet>k"
+          have "ball ?z (e / 2) \<inter> i = {}"
+          proof (clarsimp simp only: all_not_in_conv [symmetric])
+            fix y
+            assume "y \<in> ball ?z (e / 2)" and yi: "y \<in> i"
+            then have "dist ?z y < e/2" by auto
+            then have "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
+              using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
+            then have "y\<bullet>k < a\<bullet>k"
+              using e k
+              by (auto simp add: field_simps abs_less_iff as inner_simps)
+            then have "y \<notin> i"
+              unfolding ab mem_box by (auto intro!: bexI[OF _ k])
+            then show False using yi by auto
+          qed
+          moreover
+          have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)"
+            apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]])
+          proof
+            fix y
+            assume as: "y \<in> ball ?z (e/2)"
+            have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R k)"
+              apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"])
+              unfolding norm_scaleR norm_Basis[OF k]
+              apply auto
+              done
+            also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
+              apply (rule add_strict_left_mono)
+              using as e
+              apply (auto simp add: field_simps dist_norm)
+              done
+            finally show "y \<in> ball x e"
+              unfolding mem_ball dist_norm using e by (auto simp add:field_simps)
+          qed
+          ultimately show ?thesis
+            apply (rule_tac x="?z" in exI)
+            unfolding Union_insert
+            apply auto
+            done
+        next
+          let ?z = "x + (e/2) *\<^sub>R k"
+          assume as: "x\<bullet>k = b\<bullet>k"
+          have "ball ?z (e / 2) \<inter> i = {}"
+          proof (clarsimp simp only: all_not_in_conv [symmetric])
+            fix y
+            assume "y \<in> ball ?z (e / 2)" and yi: "y \<in> i"
+            then have "dist ?z y < e/2"
+              by auto
+            then have "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
+              using Basis_le_norm[OF k, of "?z - y"]
+              unfolding dist_norm by auto
+            then have "y\<bullet>k > b\<bullet>k"
+              using e k
+              by (auto simp add:field_simps inner_simps inner_Basis as)
+            then have "y \<notin> i"
+              unfolding ab mem_box by (auto intro!: bexI[OF _ k])
+            then show False using yi by auto
+          qed
+          moreover
+          have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)"
+            apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]])
+          proof
+            fix y
+            assume as: "y\<in> ball ?z (e/2)"
+            have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R k)"
+              apply (rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"])
+              unfolding norm_scaleR
+              apply (auto simp: k)
+              done
+            also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
+              apply (rule add_strict_left_mono)
+              using as unfolding mem_ball dist_norm
+              using e apply (auto simp add: field_simps)
+              done
+            finally show "y \<in> ball x e"
+              unfolding mem_ball dist_norm using e by (auto simp add:field_simps)
+          qed
+          ultimately show ?thesis
+            apply (rule_tac x="?z" in exI)
+            unfolding Union_insert
+            apply auto
+            done
+        qed
+        then obtain x where "ball x (e / 2) \<subseteq> s \<inter> \<Union>f" ..
+        then have "x \<in> s \<inter> interior (\<Union>f)"
+          unfolding lem1[where U="\<Union>f", symmetric]
+          using centre_in_ball e by auto
+        then show ?thesis
+          using insert.hyps(3) insert.prems(1) by blast
+      qed
+    qed
+  qed
+  from this[OF assms(1,3)] x
+  obtain t x e where "t \<in> f" "0 < e" "ball x e \<subseteq> s \<inter> t"
+    by blast
+  then have "x \<in> s" "x \<in> interior t"
+    using open_subset_interior[OF open_ball, of x e t]
+    by auto
+  then show False
+    using \<open>t \<in> f\<close> assms(4) by auto
+qed
+
+subsection \<open>Bounds on intervals where they exist.\<close>
+
+definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
+
+definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
+   where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
+
+lemma interval_upperbound[simp]:
+  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
+  unfolding interval_upperbound_def euclidean_representation_setsum cbox_def
+  by (safe intro!: cSup_eq) auto
+
+lemma interval_lowerbound[simp]:
+  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
+    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
+  unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def
+  by (safe intro!: cInf_eq) auto
+
+lemmas interval_bounds = interval_upperbound interval_lowerbound
+
+lemma
+  fixes X::"real set"
+  shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
+    and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
+  by (auto simp: interval_upperbound_def interval_lowerbound_def)
+
+lemma interval_bounds'[simp]:
+  assumes "cbox a b \<noteq> {}"
+  shows "interval_upperbound (cbox a b) = b"
+    and "interval_lowerbound (cbox a b) = a"
+  using assms unfolding box_ne_empty by auto
+
+
+lemma interval_upperbound_Times:
+  assumes "A \<noteq> {}" and "B \<noteq> {}"
+  shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
+proof-
+  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
+  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
+  ultimately show ?thesis unfolding interval_upperbound_def
+      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+qed
+
+lemma interval_lowerbound_Times:
+  assumes "A \<noteq> {}" and "B \<noteq> {}"
+  shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
+proof-
+  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
+  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
+  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
+      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
+  ultimately show ?thesis unfolding interval_lowerbound_def
+      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
+qed
+
+subsection \<open>Content (length, area, volume...) of an interval.\<close>
+
+definition "content (s::('a::euclidean_space) set) =
+  (if s = {} then 0 else (\<Prod>i\<in>Basis. (interval_upperbound s)\<bullet>i - (interval_lowerbound s)\<bullet>i))"
+
+lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
+  unfolding box_eq_empty unfolding not_ex not_less by auto
+
+lemma content_cbox:
+  fixes a :: "'a::euclidean_space"
+  assumes "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+  shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
+  using interval_not_empty[OF assms]
+  unfolding content_def
+  by auto
+
+lemma content_cbox':
+  fixes a :: "'a::euclidean_space"
+  assumes "cbox a b \<noteq> {}"
+  shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
+    using assms box_ne_empty(1) content_cbox by blast
+
+lemma content_real: "a \<le> b \<Longrightarrow> content {a..b} = b - a"
+  by (auto simp: interval_upperbound_def interval_lowerbound_def content_def)
+
+lemma abs_eq_content: "\<bar>y - x\<bar> = (if x\<le>y then content {x .. y} else content {y..x})"
+  by (auto simp: content_real)
+
+lemma content_singleton[simp]: "content {a} = 0"
+proof -
+  have "content (cbox a a) = 0"
+    by (subst content_cbox) (auto simp: ex_in_conv)
+  then show ?thesis by (simp add: cbox_sing)
+qed
+
+lemma content_unit[iff]: "content(cbox 0 (One::'a::euclidean_space)) = 1"
+ proof -
+   have *: "\<forall>i\<in>Basis. (0::'a)\<bullet>i \<le> (One::'a)\<bullet>i"
+    by auto
+  have "0 \<in> cbox 0 (One::'a)"
+    unfolding mem_box by auto
+  then show ?thesis
+     unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto
+ qed
+
+lemma content_pos_le[intro]:
+  fixes a::"'a::euclidean_space"
+  shows "0 \<le> content (cbox a b)"
+proof (cases "cbox a b = {}")
+  case False
+  then have *: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
+    unfolding box_ne_empty .
+  have "0 \<le> (\<Prod>i\<in>Basis. interval_upperbound (cbox a b) \<bullet> i - interval_lowerbound (cbox a b) \<bullet> i)"
+    apply (rule setprod_nonneg)
+    unfolding interval_bounds[OF *]
+    using *
+    apply auto
+    done
+  also have "\<dots> = content (cbox a b)" using False by (simp add: content_def)
+  finally show ?thesis .
+qed (simp add: content_def)
+
+corollary content_nonneg [simp]:
+  fixes a::"'a::euclidean_space"
+  shows "~ content (cbox a b) < 0"
+using not_le by blast
+
+lemma content_pos_lt:
+  fixes a :: "'a::euclidean_space"
+  assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
+  shows "0 < content (cbox a b)"
+  using assms
+  by (auto simp: content_def box_eq_empty intro!: setprod_pos)
+
+lemma content_eq_0:
+  "content (cbox a b) = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)"
+  by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI)
+
+lemma cond_cases: "(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)"
+  by auto
+
+lemma content_cbox_cases:
+  "content (cbox a (b::'a::euclidean_space)) =
+    (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis else 0)"
+  by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox)
+
+lemma content_eq_0_interior: "content (cbox a b) = 0 \<longleftrightarrow> interior(cbox a b) = {}"
+  unfolding content_eq_0 interior_cbox box_eq_empty
+  by auto
+
+lemma content_pos_lt_eq:
+  "0 < content (cbox a (b::'a::euclidean_space)) \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
+proof (rule iffI)
+  assume "0 < content (cbox a b)"
+  then have "content (cbox a b) \<noteq> 0" by auto
+  then show "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
+    unfolding content_eq_0 not_ex not_le by fastforce
+next
+  assume "\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i"
+  then show "0 < content (cbox a b)"
+    by (metis content_pos_lt)
+qed
+
+lemma content_empty [simp]: "content {} = 0"
+  unfolding content_def by auto
+
+lemma content_real_if [simp]: "content {a..b} = (if a \<le> b then b - a else 0)"
+  by (simp add: content_real)
+
+lemma content_subset:
+  assumes "cbox a b \<subseteq> cbox c d"
+  shows "content (cbox a b) \<le> content (cbox c d)"
+proof (cases "cbox a b = {}")
+  case True
+  then show ?thesis
+    using content_pos_le[of c d] by auto
+next
+  case False
+  then have ab_ne: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
+    unfolding box_ne_empty by auto
+  then have ab_ab: "a\<in>cbox a b" "b\<in>cbox a b"
+    unfolding mem_box by auto
+  have "cbox c d \<noteq> {}" using assms False by auto
+  then have cd_ne: "\<forall>i\<in>Basis. c \<bullet> i \<le> d \<bullet> i"
+    using assms unfolding box_ne_empty by auto
+  have "\<And>i. i \<in> Basis \<Longrightarrow> 0 \<le> b \<bullet> i - a \<bullet> i"
+    using ab_ne by auto
+  moreover
+  have "\<And>i. i \<in> Basis \<Longrightarrow> b \<bullet> i - a \<bullet> i \<le> d \<bullet> i - c \<bullet> i"
+    using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)]
+          assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)]
+      by (metis diff_mono)
+  ultimately show ?thesis
+    unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
+    by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \<open>cbox c d \<noteq> {}\<close>])
+qed
+
+lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0"
+  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce
+
+lemma content_times[simp]: "content (A \<times> B) = content A * content B"
+proof (cases "A \<times> B = {}")
+  let ?ub1 = "interval_upperbound" and ?lb1 = "interval_lowerbound"
+  let ?ub2 = "interval_upperbound" and ?lb2 = "interval_lowerbound"
+  assume nonempty: "A \<times> B \<noteq> {}"
+  hence "content (A \<times> B) = (\<Prod>i\<in>Basis. (?ub1 A, ?ub2 B) \<bullet> i - (?lb1 A, ?lb2 B) \<bullet> i)"
+      unfolding content_def by (simp add: interval_upperbound_Times interval_lowerbound_Times)
+  also have "... = content A * content B" unfolding content_def using nonempty
+    apply (subst Basis_prod_def, subst setprod.union_disjoint, force, force, force, simp)
+    apply (subst (1 2) setprod.reindex, auto intro: inj_onI)
+    done
+  finally show ?thesis .
+qed (auto simp: content_def)
+
+lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)"
+  by (simp add: cbox_Pair_eq)
+
+lemma content_cbox_pair_eq0_D:
+   "content (cbox (a,c) (b,d)) = 0 \<Longrightarrow> content (cbox a b) = 0 \<or> content (cbox c d) = 0"
+  by (simp add: content_Pair)
+
+lemma content_eq_0_gen:
+  fixes s :: "'a::euclidean_space set"
+  assumes "bounded s"
+  shows "content s = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. \<exists>v. \<forall>x \<in> s. x \<bullet> i = v)"  (is "_ = ?rhs")
+proof safe
+  assume "content s = 0" then show ?rhs
+    apply (clarsimp simp: ex_in_conv content_def split: if_split_asm)
+    apply (rule_tac x=a in bexI)
+    apply (rule_tac x="interval_lowerbound s \<bullet> a" in exI)
+    apply (clarsimp simp: interval_upperbound_def interval_lowerbound_def)
+    apply (drule cSUP_eq_cINF_D)
+    apply (auto simp: bounded_inner_imp_bdd_above [OF assms]  bounded_inner_imp_bdd_below [OF assms])
+    done
+next
+  fix i a
+  assume "i \<in> Basis" "\<forall>x\<in>s. x \<bullet> i = a"
+  then show "content s = 0"
+    apply (clarsimp simp: content_def)
+    apply (rule_tac x=i in bexI)
+    apply (auto simp: interval_upperbound_def interval_lowerbound_def)
+    done
+qed
+
+lemma content_0_subset_gen:
+  fixes a :: "'a::euclidean_space"
+  assumes "content t = 0" "s \<subseteq> t" "bounded t" shows "content s = 0"
+proof -
+  have "bounded s"
+    using assms by (metis bounded_subset)
+  then show ?thesis
+    using assms
+    by (auto simp: content_eq_0_gen)
+qed
+
+lemma content_0_subset: "\<lbrakk>content(cbox a b) = 0; s \<subseteq> cbox a b\<rbrakk> \<Longrightarrow> content s = 0"
+  by (simp add: content_0_subset_gen bounded_cbox)
+
+
+lemma interval_split:
+  fixes a :: "'a::euclidean_space"
+  assumes "k \<in> Basis"
+  shows
+    "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)"
+    "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b"
+  apply (rule_tac[!] set_eqI)
+  unfolding Int_iff mem_box mem_Collect_eq
+  using assms
+  apply auto
+  done
+
+lemma content_split:
+  fixes a :: "'a::euclidean_space"
+  assumes "k \<in> Basis"
+  shows "content (cbox a b) = content(cbox a b \<inter> {x. x\<bullet>k \<le> c}) + content(cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+proof cases
+  note simps = interval_split[OF assms] content_cbox_cases
+  have *: "Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}"
+    using assms by auto
+  have *: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
+    "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
+    apply (subst *(1))
+    defer
+    apply (subst *(1))
+    unfolding setprod.insert[OF *(2-)]
+    apply auto
+    done
+  assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
+  moreover
+  have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c \<Longrightarrow>
+    x * (b\<bullet>k - a\<bullet>k) = x * (max (a \<bullet> k) c - a \<bullet> k) + x * (b \<bullet> k - max (a \<bullet> k) c)"
+    by  (auto simp add: field_simps)
+  moreover
+  have **: "(\<Prod>i\<in>Basis. ((\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i - a \<bullet> i)) =
+      (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)"
+    "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) =
+      (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))"
+    by (auto intro!: setprod.cong)
+  have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
+    unfolding not_le
+    using as[unfolded ,rule_format,of k] assms
+    by auto
+  ultimately show ?thesis
+    using assms
+    unfolding simps **
+    unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"]
+    unfolding *(2)
+    by auto
+next
+  assume "\<not> (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)"
+  then have "cbox a b = {}"
+    unfolding box_eq_empty by (auto simp: not_le)
+  then show ?thesis
+    by (auto simp: not_le)
+qed
+
+subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
+
+definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
+
+lemma gaugeI:
+  assumes "\<And>x. x \<in> g x"
+    and "\<And>x. open (g x)"
+  shows "gauge g"
+  using assms unfolding gauge_def by auto
+
+lemma gaugeD[dest]:
+  assumes "gauge d"
+  shows "x \<in> d x"
+    and "open (d x)"
+  using assms unfolding gauge_def by auto
+
+lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))"
+  unfolding gauge_def by auto
+
+lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)"
+  unfolding gauge_def by auto
+
+lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
+  by (rule gauge_ball) auto
+
+lemma gauge_inter[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)"
+  unfolding gauge_def by auto
+
+lemma gauge_inters:
+  assumes "finite s"
+    and "\<forall>d\<in>s. gauge (f d)"
+  shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> s})"
+proof -
+  have *: "\<And>x. {f d x |d. d \<in> s} = (\<lambda>d. f d x) ` s"
+    by auto
+  show ?thesis
+    unfolding gauge_def unfolding *
+    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
+qed
+
+lemma gauge_existence_lemma:
+  "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
+  by (metis zero_less_one)
+
+
+subsection \<open>Divisions.\<close>
+
+definition division_of (infixl "division'_of" 40)
+where
+  "s division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>k1\<in>s. \<forall>k2\<in>s. k1 \<noteq> k2 \<longrightarrow> interior(k1) \<inter> interior(k2) = {}) \<and>
+    (\<Union>s = i)"
+
+lemma division_ofD[dest]:
+  assumes "s division_of i"
+  shows "finite s"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
+    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}"
+    and "\<Union>s = i"
+  using assms unfolding division_of_def by auto
+
+lemma division_ofI:
+  assumes "finite s"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
+    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+    and "\<Union>s = i"
+  shows "s division_of i"
+  using assms unfolding division_of_def by auto
+
+lemma division_of_finite: "s division_of i \<Longrightarrow> finite s"
+  unfolding division_of_def by auto
+
+lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)"
+  unfolding division_of_def by auto
+
+lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}"
+  unfolding division_of_def by auto
+
+lemma division_of_sing[simp]:
+  "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
+  (is "?l = ?r")
+proof
+  assume ?r
+  moreover
+  { fix k
+    assume "s = {{a}}" "k\<in>s"
+    then have "\<exists>x y. k = cbox x y"
+      apply (rule_tac x=a in exI)+
+      apply (force simp: cbox_sing)
+      done
+  }
+  ultimately show ?l
+    unfolding division_of_def cbox_sing by auto
+next
+  assume ?l
+  note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]]
+  {
+    fix x
+    assume x: "x \<in> s" have "x = {a}"
+      using *(2)[rule_format,OF x] by auto
+  }
+  moreover have "s \<noteq> {}"
+    using *(4) by auto
+  ultimately show ?r
+    unfolding cbox_sing by auto
+qed
+
+lemma elementary_empty: obtains p where "p division_of {}"
+  unfolding division_of_trivial by auto
+
+lemma elementary_interval: obtains p where "p division_of (cbox a b)"
+  by (metis division_of_trivial division_of_self)
+
+lemma division_contains: "s division_of i \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k"
+  unfolding division_of_def by auto
+
+lemma forall_in_division:
+  "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))"
+  unfolding division_of_def by fastforce
+
+lemma division_of_subset:
+  assumes "p division_of (\<Union>p)"
+    and "q \<subseteq> p"
+  shows "q division_of (\<Union>q)"
+proof (rule division_ofI)
+  note * = division_ofD[OF assms(1)]
+  show "finite q"
+    using "*"(1) assms(2) infinite_super by auto
+  {
+    fix k
+    assume "k \<in> q"
+    then have kp: "k \<in> p"
+      using assms(2) by auto
+    show "k \<subseteq> \<Union>q"
+      using \<open>k \<in> q\<close> by auto
+    show "\<exists>a b. k = cbox a b"
+      using *(4)[OF kp] by auto
+    show "k \<noteq> {}"
+      using *(3)[OF kp] by auto
+  }
+  fix k1 k2
+  assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
+  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
+    using assms(2) by auto
+  show "interior k1 \<inter> interior k2 = {}"
+    using *(5)[OF **] by auto
+qed auto
+
+lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
+  unfolding division_of_def by auto
+
+lemma division_of_content_0:
+  assumes "content (cbox a b) = 0" "d division_of (cbox a b)"
+  shows "\<forall>k\<in>d. content k = 0"
+  unfolding forall_in_division[OF assms(2)]
+  by (metis antisym_conv assms content_pos_le content_subset division_ofD(2))
+
+lemma division_inter:
+  fixes s1 s2 :: "'a::euclidean_space set"
+  assumes "p1 division_of s1"
+    and "p2 division_of s2"
+  shows "{k1 \<inter> k2 | k1 k2 .k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)"
+  (is "?A' division_of _")
+proof -
+  let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}"
+  have *: "?A' = ?A" by auto
+  show ?thesis
+    unfolding *
+  proof (rule division_ofI)
+    have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)"
+      by auto
+    moreover have "finite (p1 \<times> p2)"
+      using assms unfolding division_of_def by auto
+    ultimately show "finite ?A" by auto
+    have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
+      by auto
+    show "\<Union>?A = s1 \<inter> s2"
+      apply (rule set_eqI)
+      unfolding * and UN_iff
+      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)]
+      apply auto
+      done
+    {
+      fix k
+      assume "k \<in> ?A"
+      then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
+        by auto
+      then show "k \<noteq> {}"
+        by auto
+      show "k \<subseteq> s1 \<inter> s2"
+        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
+        unfolding k by auto
+      obtain a1 b1 where k1: "k1 = cbox a1 b1"
+        using division_ofD(4)[OF assms(1) k(2)] by blast
+      obtain a2 b2 where k2: "k2 = cbox a2 b2"
+        using division_ofD(4)[OF assms(2) k(3)] by blast
+      show "\<exists>a b. k = cbox a b"
+        unfolding k k1 k2 unfolding inter_interval by auto
+    }
+    fix k1 k2
+    assume "k1 \<in> ?A"
+    then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}"
+      by auto
+    assume "k2 \<in> ?A"
+    then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
+      by auto
+    assume "k1 \<noteq> k2"
+    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
+      unfolding k1 k2 by auto
+    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
+      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
+      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
+      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
+    show "interior k1 \<inter> interior k2 = {}"
+      unfolding k1 k2
+      apply (rule *)
+      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
+      done
+  qed
+qed
+
+lemma division_inter_1:
+  assumes "d division_of i"
+    and "cbox a (b::'a::euclidean_space) \<subseteq> i"
+  shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} division_of (cbox a b)"
+proof (cases "cbox a b = {}")
+  case True
+  show ?thesis
+    unfolding True and division_of_trivial by auto
+next
+  case False
+  have *: "cbox a b \<inter> i = cbox a b" using assms(2) by auto
+  show ?thesis
+    using division_inter[OF division_of_self[OF False] assms(1)]
+    unfolding * by auto
+qed
+
+lemma elementary_inter:
+  fixes s t :: "'a::euclidean_space set"
+  assumes "p1 division_of s"
+    and "p2 division_of t"
+  shows "\<exists>p. p division_of (s \<inter> t)"
+using assms division_inter by blast
+
+lemma elementary_inters:
+  assumes "finite f"
+    and "f \<noteq> {}"
+    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
+  shows "\<exists>p. p division_of (\<Inter>f)"
+  using assms
+proof (induct f rule: finite_induct)
+  case (insert x f)
+  show ?case
+  proof (cases "f = {}")
+    case True
+    then show ?thesis
+      unfolding True using insert by auto
+  next
+    case False
+    obtain p where "p division_of \<Inter>f"
+      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
+    moreover obtain px where "px division_of x"
+      using insert(5)[rule_format,OF insertI1] ..
+    ultimately show ?thesis
+      by (simp add: elementary_inter Inter_insert)
+  qed
+qed auto
+
+lemma division_disjoint_union:
+  assumes "p1 division_of s1"
+    and "p2 division_of s2"
+    and "interior s1 \<inter> interior s2 = {}"
+  shows "(p1 \<union> p2) division_of (s1 \<union> s2)"
+proof (rule division_ofI)
+  note d1 = division_ofD[OF assms(1)]
+  note d2 = division_ofD[OF assms(2)]
+  show "finite (p1 \<union> p2)"
+    using d1(1) d2(1) by auto
+  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
+    using d1(6) d2(6) by auto
+  {
+    fix k1 k2
+    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
+    moreover
+    let ?g="interior k1 \<inter> interior k2 = {}"
+    {
+      assume as: "k1\<in>p1" "k2\<in>p2"
+      have ?g
+        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
+        using assms(3) by blast
+    }
+    moreover
+    {
+      assume as: "k1\<in>p2" "k2\<in>p1"
+      have ?g
+        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
+        using assms(3) by blast
+    }
+    ultimately show ?g
+      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
+  }
+  fix k
+  assume k: "k \<in> p1 \<union> p2"
+  show "k \<subseteq> s1 \<union> s2"
+    using k d1(2) d2(2) by auto
+  show "k \<noteq> {}"
+    using k d1(3) d2(3) by auto
+  show "\<exists>a b. k = cbox a b"
+    using k d1(4) d2(4) by auto
+qed
+
+lemma partial_division_extend_1:
+  fixes a b c d :: "'a::euclidean_space"
+  assumes incl: "cbox c d \<subseteq> cbox a b"
+    and nonempty: "cbox c d \<noteq> {}"
+  obtains p where "p division_of (cbox a b)" "cbox c d \<in> p"
+proof
+  let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
+    cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
+  define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
+
+  show "cbox c d \<in> p"
+    unfolding p_def
+    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
+  {
+    fix i :: 'a
+    assume "i \<in> Basis"
+    with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
+      unfolding box_eq_empty subset_box by (auto simp: not_le)
+  }
+  note ord = this
+
+  show "p division_of (cbox a b)"
+  proof (rule division_ofI)
+    show "finite p"
+      unfolding p_def by (auto intro!: finite_PiE)
+    {
+      fix k
+      assume "k \<in> p"
+      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
+        by (auto simp: p_def)
+      then show "\<exists>a b. k = cbox a b"
+        by auto
+      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
+      proof (simp add: k box_eq_empty subset_box not_less, safe)
+        fix i :: 'a
+        assume i: "i \<in> Basis"
+        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+          by (auto simp: PiE_iff)
+        with i ord[of i]
+        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
+          by auto
+      qed
+      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
+        by auto
+      {
+        fix l
+        assume "l \<in> p"
+        then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
+          by (auto simp: p_def)
+        assume "l \<noteq> k"
+        have "\<exists>i\<in>Basis. f i \<noteq> g i"
+        proof (rule ccontr)
+          assume "\<not> ?thesis"
+          with f g have "f = g"
+            by (auto simp: PiE_iff extensional_def intro!: ext)
+          with \<open>l \<noteq> k\<close> show False
+            by (simp add: l k)
+        qed
+        then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
+        then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
+                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
+          using f g by (auto simp: PiE_iff)
+        with * ord[of i] show "interior l \<inter> interior k = {}"
+          by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
+      }
+      note \<open>k \<subseteq> cbox a b\<close>
+    }
+    moreover
+    {
+      fix x assume x: "x \<in> cbox a b"
+      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+      proof
+        fix i :: 'a
+        assume "i \<in> Basis"
+        with x ord[of i]
+        have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
+            (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
+          by (auto simp: cbox_def)
+        then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
+          by auto
+      qed
+      then obtain f where
+        f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
+        unfolding bchoice_iff ..
+      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
+        by auto
+      moreover from f have "x \<in> ?B (restrict f Basis)"
+        by (auto simp: mem_box)
+      ultimately have "\<exists>k\<in>p. x \<in> k"
+        unfolding p_def by blast
+    }
+    ultimately show "\<Union>p = cbox a b"
+      by auto
+  qed
+qed
+
+lemma partial_division_extend_interval:
+  assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
+  obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
+proof (cases "p = {}")
+  case True
+  obtain q where "q division_of (cbox a b)"
+    by (rule elementary_interval)
+  then show ?thesis
+    using True that by blast
+next
+  case False
+  note p = division_ofD[OF assms(1)]
+  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
+  proof
+    fix k
+    assume kp: "k \<in> p"
+    obtain c d where k: "k = cbox c d"
+      using p(4)[OF kp] by blast
+    have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}"
+      using p(2,3)[OF kp, unfolded k] using assms(2)
+      by (blast intro: order.trans)+
+    obtain q where "q division_of cbox a b" "cbox c d \<in> q"
+      by (rule partial_division_extend_1[OF *])
+    then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
+      unfolding k by auto
+  qed
+  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
+    using bchoice[OF div_cbox] by blast
+  { fix x
+    assume x: "x \<in> p"
+    have "q x division_of \<Union>q x"
+      apply (rule division_ofI)
+      using division_ofD[OF q(1)[OF x]]
+      apply auto
+      done }
+  then have "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
+    by (meson Diff_subset division_of_subset)
+  then have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
+    apply -
+    apply (rule elementary_inters [OF finite_imageI[OF p(1)]])
+    apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]])
+    done
+  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
+  have "d \<union> p division_of cbox a b"
+  proof -
+    have te: "\<And>s f t. s \<noteq> {} \<Longrightarrow> \<forall>i\<in>s. f i \<union> i = t \<Longrightarrow> t = \<Inter>(f ` s) \<union> \<Union>s" by auto
+    have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
+    proof (rule te[OF False], clarify)
+      fix i
+      assume i: "i \<in> p"
+      show "\<Union>(q i - {i}) \<union> i = cbox a b"
+        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
+    qed
+    { fix k
+      assume k: "k \<in> p"
+      have *: "\<And>u t s. t \<inter> s = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<inter> t = {}"
+        by auto
+      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior k = {}"
+      proof (rule *[OF inter_interior_unions_intervals])
+        note qk=division_ofD[OF q(1)[OF k]]
+        show "finite (q k - {k})" "open (interior k)" "\<forall>t\<in>q k - {k}. \<exists>a b. t = cbox a b"
+          using qk by auto
+        show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}"
+          using qk(5) using q(2)[OF k] by auto
+        show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q k - {k}))"
+          apply (rule interior_mono)+
+          using k
+          apply auto
+          done
+      qed } note [simp] = this
+    show "d \<union> p division_of (cbox a b)"
+      unfolding cbox_eq
+      apply (rule division_disjoint_union[OF d assms(1)])
+      apply (rule inter_interior_unions_intervals)
+      apply (rule p open_interior ballI)+
+      apply simp_all
+      done
+  qed
+  then show ?thesis
+    by (meson Un_upper2 that)
+qed
+
+lemma elementary_bounded[dest]:
+  fixes s :: "'a::euclidean_space set"
+  shows "p division_of s \<Longrightarrow> bounded s"
+  unfolding division_of_def by (metis bounded_Union bounded_cbox)
+
+lemma elementary_subset_cbox:
+  "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a (b::'a::euclidean_space)"
+  by (meson elementary_bounded bounded_subset_cbox)
+
+lemma division_union_intervals_exists:
+  fixes a b :: "'a::euclidean_space"
+  assumes "cbox a b \<noteq> {}"
+  obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
+proof (cases "cbox c d = {}")
+  case True
+  show ?thesis
+    apply (rule that[of "{}"])
+    unfolding True
+    using assms
+    apply auto
+    done
+next
+  case False
+  show ?thesis
+  proof (cases "cbox a b \<inter> cbox c d = {}")
+    case True
+    then show ?thesis
+      by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty)
+  next
+    case False
+    obtain u v where uv: "cbox a b \<inter> cbox c d = cbox u v"
+      unfolding inter_interval by auto
+    have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
+    obtain p where "p division_of cbox c d" "cbox u v \<in> p"
+      by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
+    note p = this division_ofD[OF this(1)]
+    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
+      apply (rule arg_cong[of _ _ interior])
+      using p(8) uv by auto
+    also have "\<dots> = {}"
+      unfolding interior_Int
+      apply (rule inter_interior_unions_intervals)
+      using p(6) p(7)[OF p(2)] p(3)
+      apply auto
+      done
+    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
+    have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
+      using p(8) unfolding uv[symmetric] by auto
+    have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
+    proof -
+      have "{cbox a b} division_of cbox a b"
+        by (simp add: assms division_of_self)
+      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
+        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
+    qed
+    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
+  qed
+qed
+
+lemma division_of_unions:
+  assumes "finite f"
+    and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
+    and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+  shows "\<Union>f division_of \<Union>\<Union>f"
+  using assms
+  by (auto intro!: division_ofI)
+
+lemma elementary_union_interval:
+  fixes a b :: "'a::euclidean_space"
+  assumes "p division_of \<Union>p"
+  obtains q where "q division_of (cbox a b \<union> \<Union>p)"
+proof -
+  note assm = division_ofD[OF assms]
+  have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)"
+    by auto
+  have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f"
+    by auto
+  {
+    presume "p = {} \<Longrightarrow> thesis"
+      "cbox a b = {} \<Longrightarrow> thesis"
+      "cbox a b \<noteq> {} \<Longrightarrow> interior (cbox a b) = {} \<Longrightarrow> thesis"
+      "p \<noteq> {} \<Longrightarrow> interior (cbox a b)\<noteq>{} \<Longrightarrow> cbox a b \<noteq> {} \<Longrightarrow> thesis"
+    then show thesis by auto
+  next
+    assume as: "p = {}"
+    obtain p where "p division_of (cbox a b)"
+      by (rule elementary_interval)
+    then show thesis
+      using as that by auto
+  next
+    assume as: "cbox a b = {}"
+    show thesis
+      using as assms that by auto
+  next
+    assume as: "interior (cbox a b) = {}" "cbox a b \<noteq> {}"
+    show thesis
+      apply (rule that[of "insert (cbox a b) p"],rule division_ofI)
+      unfolding finite_insert
+      apply (rule assm(1)) unfolding Union_insert
+      using assm(2-4) as
+      apply -
+      apply (fast dest: assm(5))+
+      done
+  next
+    assume as: "p \<noteq> {}" "interior (cbox a b) \<noteq> {}" "cbox a b \<noteq> {}"
+    have "\<forall>k\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
+    proof
+      fix k
+      assume kp: "k \<in> p"
+      from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast
+      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
+        by (meson as(3) division_union_intervals_exists)
+    qed
+    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
+    note q = division_ofD[OF this[rule_format]]
+    let ?D = "\<Union>{insert (cbox a b) (q k) | k. k \<in> p}"
+    show thesis
+    proof (rule that[OF division_ofI])
+      have *: "{insert (cbox a b) (q k) |k. k \<in> p} = (\<lambda>k. insert (cbox a b) (q k)) ` p"
+        by auto
+      show "finite ?D"
+        using "*" assm(1) q(1) by auto
+      show "\<Union>?D = cbox a b \<union> \<Union>p"
+        unfolding * lem1
+        unfolding lem2[OF as(1), of "cbox a b", symmetric]
+        using q(6)
+        by auto
+      fix k
+      assume k: "k \<in> ?D"
+      then show "k \<subseteq> cbox a b \<union> \<Union>p"
+        using q(2) by auto
+      show "k \<noteq> {}"
+        using q(3) k by auto
+      show "\<exists>a b. k = cbox a b"
+        using q(4) k by auto
+      fix k'
+      assume k': "k' \<in> ?D" "k \<noteq> k'"
+      obtain x where x: "k \<in> insert (cbox a b) (q x)" "x\<in>p"
+        using k by auto
+      obtain x' where x': "k'\<in>insert (cbox a b) (q x')" "x'\<in>p"
+        using k' by auto
+      show "interior k \<inter> interior k' = {}"
+      proof (cases "x = x'")
+        case True
+        show ?thesis
+          using True k' q(5) x' x by auto
+      next
+        case False
+        {
+          presume "k = cbox a b \<Longrightarrow> ?thesis"
+            and "k' = cbox a b \<Longrightarrow> ?thesis"
+            and "k \<noteq> cbox a b \<Longrightarrow> k' \<noteq> cbox a b \<Longrightarrow> ?thesis"
+          then show ?thesis by linarith
+        next
+          assume as': "k  = cbox a b"
+          show ?thesis
+            using as' k' q(5) x' by blast
+        next
+          assume as': "k' = cbox a b"
+          show ?thesis
+            using as' k'(2) q(5) x by blast
+        }
+        assume as': "k \<noteq> cbox a b" "k' \<noteq> cbox a b"
+        obtain c d where k: "k = cbox c d"
+          using q(4)[OF x(2,1)] by blast
+        have "interior k \<inter> interior (cbox a b) = {}"
+          using as' k'(2) q(5) x by blast
+        then have "interior k \<subseteq> interior x"
+        using interior_subset_union_intervals
+          by (metis as(2) k q(2) x interior_subset_union_intervals)
+        moreover
+        obtain c d where c_d: "k' = cbox c d"
+          using q(4)[OF x'(2,1)] by blast
+        have "interior k' \<inter> interior (cbox a b) = {}"
+          using as'(2) q(5) x' by blast
+        then have "interior k' \<subseteq> interior x'"
+          by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2))
+        ultimately show ?thesis
+          using assm(5)[OF x(2) x'(2) False] by auto
+      qed
+    qed
+  }
+qed
+
+lemma elementary_unions_intervals:
+  assumes fin: "finite f"
+    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
+  obtains p where "p division_of (\<Union>f)"
+proof -
+  have "\<exists>p. p division_of (\<Union>f)"
+  proof (induct_tac f rule:finite_subset_induct)
+    show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
+  next
+    fix x F
+    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
+    from this(3) obtain p where p: "p division_of \<Union>F" ..
+    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
+    have *: "\<Union>F = \<Union>p"
+      using division_ofD[OF p] by auto
+    show "\<exists>p. p division_of \<Union>insert x F"
+      using elementary_union_interval[OF p[unfolded *], of a b]
+      unfolding Union_insert x * by metis
+  qed (insert assms, auto)
+  then show ?thesis
+    using that by auto
+qed
+
+lemma elementary_union:
+  fixes s t :: "'a::euclidean_space set"
+  assumes "ps division_of s" "pt division_of t"
+  obtains p where "p division_of (s \<union> t)"
+proof -
+  have *: "s \<union> t = \<Union>ps \<union> \<Union>pt"
+    using assms unfolding division_of_def by auto
+  show ?thesis
+    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
+    using assms apply auto
+    by (simp add: * that)
+qed
+
+lemma partial_division_extend:
+  fixes t :: "'a::euclidean_space set"
+  assumes "p division_of s"
+    and "q division_of t"
+    and "s \<subseteq> t"
+  obtains r where "p \<subseteq> r" and "r division_of t"
+proof -
+  note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
+  obtain a b where ab: "t \<subseteq> cbox a b"
+    using elementary_subset_cbox[OF assms(2)] by auto
+  obtain r1 where "p \<subseteq> r1" "r1 division_of (cbox a b)"
+    using assms
+    by (metis ab dual_order.trans partial_division_extend_interval divp(6))
+  note r1 = this division_ofD[OF this(2)]
+  obtain p' where "p' division_of \<Union>(r1 - p)"
+    apply (rule elementary_unions_intervals[of "r1 - p"])
+    using r1(3,6)
+    apply auto
+    done
+  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
+    by (metis assms(2) divq(6) elementary_inter)
+  {
+    fix x
+    assume x: "x \<in> t" "x \<notin> s"
+    then have "x\<in>\<Union>r1"
+      unfolding r1 using ab by auto
+    then obtain r where r: "r \<in> r1" "x \<in> r"
+      unfolding Union_iff ..
+    moreover
+    have "r \<notin> p"
+    proof
+      assume "r \<in> p"
+      then have "x \<in> s" using divp(2) r by auto
+      then show False using x by auto
+    qed
+    ultimately have "x\<in>\<Union>(r1 - p)" by auto
+  }
+  then have *: "t = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
+    unfolding divp divq using assms(3) by auto
+  show ?thesis
+    apply (rule that[of "p \<union> r2"])
+    unfolding *
+    defer
+    apply (rule division_disjoint_union)
+    unfolding divp(6)
+    apply(rule assms r2)+
+  proof -
+    have "interior s \<inter> interior (\<Union>(r1-p)) = {}"
+    proof (rule inter_interior_unions_intervals)
+      show "finite (r1 - p)" and "open (interior s)" and "\<forall>t\<in>r1-p. \<exists>a b. t = cbox a b"
+        using r1 by auto
+      have *: "\<And>s. (\<And>x. x \<in> s \<Longrightarrow> False) \<Longrightarrow> s = {}"
+        by auto
+      show "\<forall>t\<in>r1-p. interior s \<inter> interior t = {}"
+      proof
+        fix m x
+        assume as: "m \<in> r1 - p"
+        have "interior m \<inter> interior (\<Union>p) = {}"
+        proof (rule inter_interior_unions_intervals)
+          show "finite p" and "open (interior m)" and "\<forall>t\<in>p. \<exists>a b. t = cbox a b"
+            using divp by auto
+          show "\<forall>t\<in>p. interior m \<inter> interior t = {}"
+            by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp)
+        qed
+        then show "interior s \<inter> interior m = {}"
+          unfolding divp by auto
+      qed
+    qed
+    then show "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
+      using interior_subset by auto
+  qed auto
+qed
+
+lemma division_split_left_inj:
+  fixes type :: "'a::euclidean_space"
+  assumes "d division_of i"
+    and "k1 \<in> d"
+    and "k2 \<in> d"
+    and "k1 \<noteq> k2"
+    and "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
+    and k: "k\<in>Basis"
+  shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
+proof -
+  note d=division_ofD[OF assms(1)]
+  have *: "\<And>(a::'a) b c. content (cbox a b \<inter> {x. x\<bullet>k \<le> c}) = 0 \<longleftrightarrow>
+    interior(cbox a b \<inter> {x. x\<bullet>k \<le> c}) = {}"
+    unfolding  interval_split[OF k] content_eq_0_interior by auto
+  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
+  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
+    by auto
+  show ?thesis
+    unfolding uv1 uv2 *
+    apply (rule **[OF d(5)[OF assms(2-4)]])
+    apply (simp add: uv1)
+    using assms(5) uv1 by auto
+qed
+
+lemma division_split_right_inj:
+  fixes type :: "'a::euclidean_space"
+  assumes "d division_of i"
+    and "k1 \<in> d"
+    and "k2 \<in> d"
+    and "k1 \<noteq> k2"
+    and "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
+    and k: "k \<in> Basis"
+  shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
+proof -
+  note d=division_ofD[OF assms(1)]
+  have *: "\<And>a b::'a. \<And>c. content(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = 0 \<longleftrightarrow>
+    interior(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = {}"
+    unfolding interval_split[OF k] content_eq_0_interior by auto
+  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
+  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
+  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
+    by auto
+  show ?thesis
+    unfolding uv1 uv2 *
+    apply (rule **[OF d(5)[OF assms(2-4)]])
+    apply (simp add: uv1)
+    using assms(5) uv1 by auto
+qed
+
+
+lemma division_split:
+  fixes a :: "'a::euclidean_space"
+  assumes "p division_of (cbox a b)"
+    and k: "k\<in>Basis"
+  shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
+      (is "?p1 division_of ?I1")
+    and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+      (is "?p2 division_of ?I2")
+proof (rule_tac[!] division_ofI)
+  note p = division_ofD[OF assms(1)]
+  show "finite ?p1" "finite ?p2"
+    using p(1) by auto
+  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
+    unfolding p(6)[symmetric] by auto
+  {
+    fix k
+    assume "k \<in> ?p1"
+    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
+    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
+    show "k \<subseteq> ?I1"
+      using l p(2) uv by force
+    show  "k \<noteq> {}"
+      by (simp add: l)
+    show  "\<exists>a b. k = cbox a b"
+      apply (simp add: l uv p(2-3)[OF l(2)])
+      apply (subst interval_split[OF k])
+      apply (auto intro: order.trans)
+      done
+    fix k'
+    assume "k' \<in> ?p1"
+    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
+    assume "k \<noteq> k'"
+    then show "interior k \<inter> interior k' = {}"
+      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
+  }
+  {
+    fix k
+    assume "k \<in> ?p2"
+    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
+    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
+    show "k \<subseteq> ?I2"
+      using l p(2) uv by force
+    show  "k \<noteq> {}"
+      by (simp add: l)
+    show  "\<exists>a b. k = cbox a b"
+      apply (simp add: l uv p(2-3)[OF l(2)])
+      apply (subst interval_split[OF k])
+      apply (auto intro: order.trans)
+      done
+    fix k'
+    assume "k' \<in> ?p2"
+    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
+    assume "k \<noteq> k'"
+    then show "interior k \<inter> interior k' = {}"
+      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
+  }
+qed
+
+subsection \<open>Tagged (partial) divisions.\<close>
+
+definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
+  where "s tagged_partial_division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
+      interior k1 \<inter> interior k2 = {})"
+
+lemma tagged_partial_division_ofD[dest]:
+  assumes "s tagged_partial_division_of i"
+  shows "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow>
+      (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
+  using assms unfolding tagged_partial_division_of_def by blast+
+
+definition tagged_division_of (infixr "tagged'_division'_of" 40)
+  where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+
+lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
+  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
+
+lemma tagged_division_of:
+  "s tagged_division_of i \<longleftrightarrow>
+    finite s \<and>
+    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
+    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
+      interior k1 \<inter> interior k2 = {}) \<and>
+    (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
+
+lemma tagged_division_ofI:
+  assumes "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
+      interior k1 \<inter> interior k2 = {}"
+    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  shows "s tagged_division_of i"
+  unfolding tagged_division_of
+  using assms
+  apply auto
+  apply fastforce+
+  done
+
+lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
+  assumes "s tagged_division_of i"
+  shows "finite s"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
+    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
+    and "\<And>x1 k1 x2 k2. (x1, k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
+      interior k1 \<inter> interior k2 = {}"
+    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
+  using assms unfolding tagged_division_of by blast+
+
+lemma division_of_tagged_division:
+  assumes "s tagged_division_of i"
+  shows "(snd ` s) division_of i"
+proof (rule division_ofI)
+  note assm = tagged_division_ofD[OF assms]
+  show "\<Union>(snd ` s) = i" "finite (snd ` s)"
+    using assm by auto
+  fix k
+  assume k: "k \<in> snd ` s"
+  then obtain xk where xk: "(xk, k) \<in> s"
+    by auto
+  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
+    using assm by fastforce+
+  fix k'
+  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
+  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
+    by auto
+  then show "interior k \<inter> interior k' = {}"
+    using assm(5) k'(2) xk by blast
+qed
+
+lemma partial_division_of_tagged_division:
+  assumes "s tagged_partial_division_of i"
+  shows "(snd ` s) division_of \<Union>(snd ` s)"
+proof (rule division_ofI)
+  note assm = tagged_partial_division_ofD[OF assms]
+  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
+    using assm by auto
+  fix k
+  assume k: "k \<in> snd ` s"
+  then obtain xk where xk: "(xk, k) \<in> s"
+    by auto
+  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
+    using assm by auto
+  fix k'
+  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
+  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
+    by auto
+  then show "interior k \<inter> interior k' = {}"
+    using assm(5) k'(2) xk by auto
+qed
+
+lemma tagged_partial_division_subset:
+  assumes "s tagged_partial_division_of i"
+    and "t \<subseteq> s"
+  shows "t tagged_partial_division_of i"
+  using assms
+  unfolding tagged_partial_division_of_def
+  using finite_subset[OF assms(2)]
+  by blast
+
+lemma (in comm_monoid_set) over_tagged_division_lemma:
+  assumes "p tagged_division_of i"
+    and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> content (cbox u v) = 0 \<Longrightarrow> d (cbox u v) = \<^bold>1"
+  shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)"
+proof -
+  have *: "(\<lambda>(x,k). d k) = d \<circ> snd"
+    unfolding o_def by (rule ext) auto
+  note assm = tagged_division_ofD[OF assms(1)]
+  show ?thesis
+    unfolding *
+  proof (rule reindex_nontrivial[symmetric])
+    show "finite p"
+      using assm by auto
+    fix x y
+    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
+    obtain a b where ab: "snd x = cbox a b"
+      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
+    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
+      by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
+    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
+      by (intro assm(5)[of "fst x" _ "fst y"]) auto
+    then have "content (cbox a b) = 0"
+      unfolding \<open>snd x = snd y\<close>[symmetric] ab content_eq_0_interior by auto
+    then have "d (cbox a b) = \<^bold>1"
+      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
+    then show "d (snd x) = \<^bold>1"
+      unfolding ab by auto
+  qed
+qed
+
+lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> i"
+  by auto
+
+lemma tagged_division_of_empty: "{} tagged_division_of {}"
+  unfolding tagged_division_of by auto
+
+lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \<longleftrightarrow> p = {}"
+  unfolding tagged_partial_division_of_def by auto
+
+lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}"
+  unfolding tagged_division_of by auto
+
+lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)"
+  by (rule tagged_division_ofI) auto
+
+lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
+  unfolding box_real[symmetric]
+  by (rule tagged_division_of_self)
+
+lemma tagged_division_union:
+  assumes "p1 tagged_division_of s1"
+    and "p2 tagged_division_of s2"
+    and "interior s1 \<inter> interior s2 = {}"
+  shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
+proof (rule tagged_division_ofI)
+  note p1 = tagged_division_ofD[OF assms(1)]
+  note p2 = tagged_division_ofD[OF assms(2)]
+  show "finite (p1 \<union> p2)"
+    using p1(1) p2(1) by auto
+  show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
+    using p1(6) p2(6) by blast
+  fix x k
+  assume xk: "(x, k) \<in> p1 \<union> p2"
+  show "x \<in> k" "\<exists>a b. k = cbox a b"
+    using xk p1(2,4) p2(2,4) by auto
+  show "k \<subseteq> s1 \<union> s2"
+    using xk p1(3) p2(3) by blast
+  fix x' k'
+  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
+  have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
+    using assms(3) interior_mono by blast
+  show "interior k \<inter> interior k' = {}"
+    apply (cases "(x, k) \<in> p1")
+    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
+    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
+qed
+
+lemma tagged_division_unions:
+  assumes "finite iset"
+    and "\<forall>i\<in>iset. pfn i tagged_division_of i"
+    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}"
+  shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
+proof (rule tagged_division_ofI)
+  note assm = tagged_division_ofD[OF assms(2)[rule_format]]
+  show "finite (\<Union>(pfn ` iset))"
+    apply (rule finite_Union)
+    using assms
+    apply auto
+    done
+  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
+    by blast
+  also have "\<dots> = \<Union>iset"
+    using assm(6) by auto
+  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" .
+  fix x k
+  assume xk: "(x, k) \<in> \<Union>(pfn ` iset)"
+  then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i"
+    by auto
+  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset"
+    using assm(2-4)[OF i] using i(1) by auto
+  fix x' k'
+  assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')"
+  then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'"
+    by auto
+  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
+    using i(1) i'(1)
+    using assms(3)[rule_format] interior_mono
+    by blast
+  show "interior k \<inter> interior k' = {}"
+    apply (cases "i = i'")
+    using assm(5) i' i(2) xk'(2) apply blast
+    using "*" assm(3) i' i by auto
+qed
+
+lemma tagged_partial_division_of_union_self:
+  assumes "p tagged_partial_division_of s"
+  shows "p tagged_division_of (\<Union>(snd ` p))"
+  apply (rule tagged_division_ofI)
+  using tagged_partial_division_ofD[OF assms]
+  apply auto
+  done
+
+lemma tagged_division_of_union_self:
+  assumes "p tagged_division_of s"
+  shows "p tagged_division_of (\<Union>(snd ` p))"
+  apply (rule tagged_division_ofI)
+  using tagged_division_ofD[OF assms]
+  apply auto
+  done
+
+subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close>
+
+text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is
+  @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and
+  @{typ bool}.\<close>
+
+lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
+  using content_empty unfolding empty_as_interval by auto
+
+paragraph \<open>Using additivity of lifted function to encode definedness.\<close>
+
+definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
+where
+  "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
+
+lemma lift_option_simps[simp]:
+  "lift_option f (Some a) (Some b) = Some (f a b)"
+  "lift_option f None b' = None"
+  "lift_option f a' None = None"
+  by (auto simp: lift_option_def)
+
+lemma (in comm_monoid) comm_monoid_lift_option: "comm_monoid (lift_option op \<^bold>*) (Some \<^bold>1)"
+  proof qed (auto simp: lift_option_def ac_simps split: bind_split)
+
+lemma (in comm_monoid) comm_monoid_set_lift_option: "comm_monoid_set (lift_option op \<^bold>*) (Some \<^bold>1)"
+  proof qed (auto simp: lift_option_def ac_simps split: bind_split)
+
+lemma comm_monoid_and: "comm_monoid op \<and> True"
+  proof qed auto
+
+lemma comm_monoid_set_and: "comm_monoid_set op \<and> True"
+  proof qed auto
+
+paragraph \<open>Operative\<close>
+
+definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool"
+  where "operative g \<longleftrightarrow>
+    (\<forall>a b. content (cbox a b) = 0 \<longrightarrow> g (cbox a b) = \<^bold>1) \<and>
+    (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))"
+
+lemma (in comm_monoid) operativeD[dest]:
+  assumes "operative g"
+  shows "\<And>a b. content (cbox a b) = 0 \<Longrightarrow> g (cbox a b) = \<^bold>1"
+    and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+  using assms unfolding operative_def by auto
+
+lemma (in comm_monoid) operative_empty: "operative g \<Longrightarrow> g {} = \<^bold>1"
+  unfolding operative_def by (rule property_empty_interval) auto
+
+lemma operative_content[intro]: "add.operative content"
+  by (force simp add: add.operative_def content_split[symmetric])
+
+definition "division_points (k::('a::euclidean_space) set) d =
+   {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
+     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
+
+lemma division_points_finite:
+  fixes i :: "'a::euclidean_space set"
+  assumes "d division_of i"
+  shows "finite (division_points i d)"
+proof -
+  note assm = division_ofD[OF assms]
+  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
+    (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
+  have *: "division_points i d = \<Union>(?M ` Basis)"
+    unfolding division_points_def by auto
+  show ?thesis
+    unfolding * using assm by auto
+qed
+
+lemma division_points_subset:
+  fixes a :: "'a::euclidean_space"
+  assumes "d division_of (cbox a b)"
+    and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
+    and k: "k \<in> Basis"
+  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
+      division_points (cbox a b) d" (is ?t1)
+    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
+      division_points (cbox a b) d" (is ?t2)
+proof -
+  note assm = division_ofD[OF assms(1)]
+  have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+    "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
+    "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
+    "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
+    using assms using less_imp_le by auto
+  show ?t1 (*FIXME a horrible mess*)
+    unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
+    unfolding *
+    apply (rule subsetI)
+    unfolding mem_Collect_eq split_beta
+    apply (erule bexE conjE)+
+    apply (simp add: )
+    apply (erule exE conjE)+
+  proof
+    fix i l x
+    assume as:
+      "a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)"
+      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
+      and fstx: "fst x \<in> Basis"
+    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
+    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
+      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
+    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+      using l using as(6) unfolding box_ne_empty[symmetric] by auto
+    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
+      using as(1-3,5) fstx
+      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+      apply (auto split: if_split_asm)
+      done
+    show "snd x < b \<bullet> fst x"
+      using as(2) \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
+  qed
+  show ?t2
+    unfolding division_points_def interval_split[OF k, of a b]
+    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
+    unfolding *
+    unfolding subset_eq
+    apply rule
+    unfolding mem_Collect_eq split_beta
+    apply (erule bexE conjE)+
+    apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
+    apply (erule exE conjE)+
+  proof
+    fix i l x
+    assume as:
+      "(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x"
+      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
+      and fstx: "fst x \<in> Basis"
+    from assm(4)[OF this(5)] guess u v by (elim exE) note l=this
+    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
+      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
+    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
+      using l using as(6) unfolding box_ne_empty[symmetric] by auto
+    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
+      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
+      using as(1-3,5) fstx
+      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
+      apply (auto split: if_split_asm)
+      done
+    show "a \<bullet> fst x < snd x"
+      using as(1) \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
+   qed
+qed
+
+lemma division_points_psubset:
+  fixes a :: "'a::euclidean_space"
+  assumes "d division_of (cbox a b)"
+      and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
+      and "l \<in> d"
+      and "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c"
+      and k: "k \<in> Basis"
+  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset>
+         division_points (cbox a b) d" (is "?D1 \<subset> ?D")
+    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
+         division_points (cbox a b) d" (is "?D2 \<subset> ?D")
+proof -
+  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+    using assms(2) by (auto intro!:less_imp_le)
+  guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this
+  have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
+    using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty
+    using subset_box(1)
+    apply auto
+    apply blast+
+    done
+  have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
+          "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
+    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
+    using uv[rule_format, of k] ab k
+    by auto
+  have "\<exists>x. x \<in> ?D - ?D1"
+    using assms(3-)
+    unfolding division_points_def interval_bounds[OF ab]
+    apply -
+    apply (erule disjE)
+    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
+    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
+    done
+  moreover have "?D1 \<subseteq> ?D"
+    by (auto simp add: assms division_points_subset)
+  ultimately show "?D1 \<subset> ?D"
+    by blast
+  have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
+    "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
+    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
+    using uv[rule_format, of k] ab k
+    by auto
+  have "\<exists>x. x \<in> ?D - ?D2"
+    using assms(3-)
+    unfolding division_points_def interval_bounds[OF ab]
+    apply -
+    apply (erule disjE)
+    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
+    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
+    done
+  moreover have "?D2 \<subseteq> ?D"
+    by (auto simp add: assms division_points_subset)
+  ultimately show "?D2 \<subset> ?D"
+    by blast
+qed
+
+lemma (in comm_monoid_set) operative_division:
+  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
+  assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)"
+proof -
+  define C where [abs_def]: "C = card (division_points (cbox a b) d)"
+  then show ?thesis
+    using d
+  proof (induction C arbitrary: a b d rule: less_induct)
+    case (less a b d)
+    show ?case
+    proof cases
+      show "content (cbox a b) = 0 \<Longrightarrow> F g d = g (cbox a b)"
+        using division_of_content_0[OF _ less.prems] operativeD(1)[OF  g] division_ofD(4)[OF less.prems]
+        by (fastforce intro!: neutral)
+    next
+      assume "content (cbox a b) \<noteq> 0"
+      note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq]
+      then have ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
+        by (auto intro!: less_imp_le)
+      show "F g d = g (cbox a b)"
+      proof (cases "division_points (cbox a b) d = {}")
+        case True
+        { fix u v and j :: 'b
+          assume j: "j \<in> Basis" and as: "cbox u v \<in> d"
+          then have "cbox u v \<noteq> {}"
+            using less.prems by blast
+          then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j"
+            using j unfolding box_ne_empty by auto
+          have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)"
+            using as j by auto
+          have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d"
+               "(j, v\<bullet>j) \<notin> division_points (cbox a b) d" using True by auto
+          note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
+          note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
+          moreover
+          have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
+            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
+            apply (metis j subset_box(1) uv(1))
+            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
+          ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
+            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
+        then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
+          (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
+          unfolding forall_in_division[OF less.prems] by blast
+        have "(1/2) *\<^sub>R (a+b) \<in> cbox a b"
+          unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps)
+        note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
+        then guess i .. note i=this
+        guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this
+        have "cbox a b \<in> d"
+        proof -
+          have "u = a" "v = b"
+            unfolding euclidean_eq_iff[where 'a='b]
+          proof safe
+            fix j :: 'b
+            assume j: "j \<in> Basis"
+            note i(2)[unfolded uv mem_box,rule_format,of j]
+            then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
+              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
+          qed
+          then have "i = cbox a b" using uv by auto
+          then show ?thesis using i by auto
+        qed
+        then have deq: "d = insert (cbox a b) (d - {cbox a b})"
+          by auto
+        have "F g (d - {cbox a b}) = \<^bold>1"
+        proof (intro neutral ballI)
+          fix x
+          assume x: "x \<in> d - {cbox a b}"
+          then have "x\<in>d"
+            by auto note d'[rule_format,OF this]
+          then guess u v by (elim exE conjE) note uv=this
+          have "u \<noteq> a \<or> v \<noteq> b"
+            using x[unfolded uv] by auto
+          then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis"
+            unfolding euclidean_eq_iff[where 'a='b] by auto
+          then have "u\<bullet>j = v\<bullet>j"
+            using uv(2)[rule_format,OF j] by auto
+          then have "content (cbox u v) = 0"
+            unfolding content_eq_0 using j
+            by force
+          then show "g x = \<^bold>1"
+            unfolding uv(1) by (rule operativeD(1)[OF g])
+        qed
+        then show "F g d = g (cbox a b)"
+          using division_ofD[OF less.prems]
+          apply (subst deq)
+          apply (subst insert)
+          apply auto
+          done
+      next
+        case False
+        then have "\<exists>x. x \<in> division_points (cbox a b) d"
+          by auto
+        then guess k c
+          unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv
+          apply (elim exE conjE)
+          done
+        note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
+        from this(3) guess j .. note j=this
+        define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+        define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+        define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
+        define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
+        note division_points_psubset[OF \<open>d division_of cbox a b\<close> ab kc(1-2) j]
+        note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
+        then have *: "F g d1 = g (cbox a b \<inter> {x. x\<bullet>k \<le> c})" "F g d2 = g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+          unfolding interval_split[OF kc(4)]
+          apply (rule_tac[!] "less.hyps"[rule_format])
+          using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c]
+          apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
+          done
+        { fix l y
+          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y"
+          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
+          have "g (l \<inter> {x. x \<bullet> k \<le> c}) = \<^bold>1"
+            unfolding leq interval_split[OF kc(4)]
+            apply (rule operativeD[OF g])
+            unfolding interval_split[symmetric, OF kc(4)]
+            using division_split_left_inj less as kc leq by blast
+        } note fxk_le = this
+        { fix l y
+          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y"
+          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
+          have "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
+            unfolding leq interval_split[OF kc(4)]
+            apply (rule operativeD(1)[OF g])
+            unfolding interval_split[symmetric,OF kc(4)]
+            using division_split_right_inj less leq as kc by blast
+        } note fxk_ge = this
+        have d1_alt: "d1 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<le> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+          using d1_def by auto
+        have d2_alt: "d2 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<ge> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+          using d2_def by auto
+        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
+          unfolding * using g kc(4) by blast
+        also have "F g d1 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d"
+          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
+          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
+        also have "F g d2 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d"
+          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
+          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
+        also have *: "\<forall>x\<in>d. g x = g (x \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (x \<inter> {x. c \<le> x \<bullet> k})"
+          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
+          using g kc(4) by blast
+        have "F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d \<^bold>* F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d = F g d"
+          using * by (simp add: distrib)
+        finally show ?thesis by auto
+      qed
+    qed
+  qed
+qed
+
+lemma (in comm_monoid_set) operative_tagged_division:
+  assumes f: "operative g" and d: "d tagged_division_of (cbox a b)"
+  shows "F (\<lambda>(x, l). g l) d = g (cbox a b)"
+  unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric]
+  by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d])
+
+lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> setsum content d = content (cbox a b)"
+  by (metis operative_content setsum.operative_division)
+
+lemma additive_content_tagged_division:
+  "d tagged_division_of (cbox a b) \<Longrightarrow> setsum (\<lambda>(x,l). content l) d = content (cbox a b)"
+  unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast
+
+lemma
+  shows real_inner_1_left: "inner 1 x = x"
+  and real_inner_1_right: "inner x 1 = x"
+  by simp_all
+
+lemma content_real_eq_0: "content {a .. b::real} = 0 \<longleftrightarrow> a \<ge> b"
+  by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0)
+
+lemma interval_real_split:
+  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
+  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
+  apply (metis Int_atLeastAtMostL1 atMost_def)
+  apply (metis Int_atLeastAtMostL2 atLeast_def)
+  done
+
+lemma (in comm_monoid) operative_1_lt:
+  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
+    ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
+  apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric]
+              del: content_real_if)
+proof safe
+  fix a b c :: real
+  assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
+  assume "a < c" "c < b"
+  with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    by (simp add: less_imp_le min.absorb2 max.absorb2)
+next
+  fix a b c :: real
+  assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1"
+    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+  from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2)
+  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
+    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    by auto
+  show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
+    by (auto simp: min_def max_def le_less)
+qed
+
+lemma (in comm_monoid) operative_1_le:
+  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
+    ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
+  unfolding operative_1_lt
+proof safe
+  fix a b c :: real
+  assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b"
+  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    apply (rule as(1)[rule_format])
+    using as(2-)
+    apply auto
+    done
+next
+  fix a b c :: real
+  assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1"
+    and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
+    and "a \<le> c"
+    and "c \<le> b"
+  note as = this[rule_format]
+  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
+  proof (cases "c = a \<or> c = b")
+    case False
+    then show ?thesis
+      apply -
+      apply (subst as(2))
+      using as(3-)
+      apply auto
+      done
+  next
+    case True
+    then show ?thesis
+    proof
+      assume *: "c = a"
+      then have "g {a .. c} = \<^bold>1"
+        apply -
+        apply (rule as(1)[rule_format])
+        apply auto
+        done
+      then show ?thesis
+        unfolding * by auto
+    next
+      assume *: "c = b"
+      then have "g {c .. b} = \<^bold>1"
+        apply -
+        apply (rule as(1)[rule_format])
+        apply auto
+        done
+      then show ?thesis
+        unfolding * by auto
+    qed
+  qed
+qed
+
+subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
+
+definition fine  (infixr "fine" 46)
+  where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
+
+lemma fineI:
+  assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x"
+  shows "d fine s"
+  using assms unfolding fine_def by auto
+
+lemma fineD[dest]:
+  assumes "d fine s"
+  shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x"
+  using assms unfolding fine_def by auto
+
+lemma fine_inter: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
+  unfolding fine_def by auto
+
+lemma fine_inters:
+ "(\<lambda>x. \<Inter>{f d x | d.  d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)"
+  unfolding fine_def by blast
+
+lemma fine_union: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
+  unfolding fine_def by blast
+
+lemma fine_unions: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
+  unfolding fine_def by auto
+
+lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
+  unfolding fine_def by blast
+
+
+subsection \<open>Gauge integral. Define on compact intervals first, then use a limit.\<close>
+
+definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46)
+  where "(f has_integral_compact_interval y) i \<longleftrightarrow>
+    (\<forall>e>0. \<exists>d. gauge d \<and>
+      (\<forall>p. p tagged_division_of i \<and> d fine p \<longrightarrow>
+        norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
+
+definition has_integral ::
+    "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool"
+  (infixr "has'_integral" 46)
+  where "(f has_integral y) i \<longleftrightarrow>
+    (if \<exists>a b. i = cbox a b
+     then (f has_integral_compact_interval y) i
+     else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) (cbox a b) \<and>
+        norm (z - y) < e)))"
+
+lemma has_integral:
+  "(f has_integral y) (cbox a b) \<longleftrightarrow>
+    (\<forall>e>0. \<exists>d. gauge d \<and>
+      (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+        norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
+  unfolding has_integral_def has_integral_compact_interval_def
+  by auto
+
+lemma has_integral_real:
+  "(f has_integral y) {a .. b::real} \<longleftrightarrow>
+    (\<forall>e>0. \<exists>d. gauge d \<and>
+      (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
+        norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
+  unfolding box_real[symmetric]
+  by (rule has_integral)
+
+lemma has_integralD[dest]:
+  assumes "(f has_integral y) (cbox a b)"
+    and "e > 0"
+  obtains d where "gauge d"
+    and "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d fine p \<Longrightarrow>
+      norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f(x)) p - y) < e"
+  using assms unfolding has_integral by auto
+
+lemma has_integral_alt:
+  "(f has_integral y) i \<longleftrightarrow>
+    (if \<exists>a b. i = cbox a b
+     then (f has_integral y) i
+     else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e)))"
+  unfolding has_integral
+  unfolding has_integral_compact_interval_def has_integral_def
+  by auto
+
+lemma has_integral_altD:
+  assumes "(f has_integral y) i"
+    and "\<not> (\<exists>a b. i = cbox a b)"
+    and "e>0"
+  obtains B where "B > 0"
+    and "\<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - y) < e)"
+  using assms
+  unfolding has_integral
+  unfolding has_integral_compact_interval_def has_integral_def
+  by auto
+
+definition integrable_on (infixr "integrable'_on" 46)
+  where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)"
+
+definition "integral i f = (SOME y. (f has_integral y) i \<or> ~ f integrable_on i \<and> y=0)"
+
+lemma integrable_integral[dest]: "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i"
+  unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex)
+
+lemma not_integrable_integral: "~ f integrable_on i \<Longrightarrow> integral i f = 0"
+  unfolding integrable_on_def integral_def by blast
+
+lemma has_integral_integrable[intro]: "(f has_integral i) s \<Longrightarrow> f integrable_on s"
+  unfolding integrable_on_def by auto
+
+lemma has_integral_integral: "f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
+  by auto
+
+lemma setsum_content_null:
+  assumes "content (cbox a b) = 0"
+    and "p tagged_division_of (cbox a b)"
+  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
+proof (rule setsum.neutral, rule)
+  fix y
+  assume y: "y \<in> p"
+  obtain x k where xk: "y = (x, k)"
+    using surj_pair[of y] by blast
+  note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]]
+  from this(2) obtain c d where k: "k = cbox c d" by blast
+  have "(\<lambda>(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x"
+    unfolding xk by auto
+  also have "\<dots> = 0"
+    using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d]
+    unfolding assms(1) k
+    by auto
+  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
+qed
+
+
+subsection \<open>Some basic combining lemmas.\<close>
+
+lemma tagged_division_unions_exists:
+  assumes "finite iset"
+    and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p"
+    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
+    and "\<Union>iset = i"
+   obtains p where "p tagged_division_of i" and "d fine p"
+proof -
+  obtain pfn where pfn:
+    "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x"
+    "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x"
+    using bchoice[OF assms(2)] by auto
+  show thesis
+    apply (rule_tac p="\<Union>(pfn ` iset)" in that)
+    using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force
+    by (metis (mono_tags, lifting) fine_unions imageE pfn(2))
+qed
+
+
+subsection \<open>The set we're concerned with must be closed.\<close>
+
+lemma division_of_closed:
+  fixes i :: "'n::euclidean_space set"
+  shows "s division_of i \<Longrightarrow> closed i"
+  unfolding division_of_def by fastforce
+
+subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
+
+lemma interval_bisection_step:
+  fixes type :: "'a::euclidean_space"
+  assumes "P {}"
+    and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)"
+    and "\<not> P (cbox a (b::'a))"
+  obtains c d where "\<not> P (cbox c d)"
+    and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
+proof -
+  have "cbox a b \<noteq> {}"
+    using assms(1,3) by metis
+  then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
+    by (force simp: mem_box)
+  { fix f
+    have "\<lbrakk>finite f;
+           \<And>s. s\<in>f \<Longrightarrow> P s;
+           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
+           \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)"
+    proof (induct f rule: finite_induct)
+      case empty
+      show ?case
+        using assms(1) by auto
+    next
+      case (insert x f)
+      show ?case
+        unfolding Union_insert
+        apply (rule assms(2)[rule_format])
+        using inter_interior_unions_intervals [of f "interior x"]
+        apply (auto simp: insert)
+        by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff)
+    qed
+  } note UN_cases = this
+  let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or>
+    (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
+  let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
+  {
+    presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False"
+    then show thesis
+      unfolding atomize_not not_all
+      by (blast intro: that)
+  }
+  assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)"
+  have "P (\<Union>?A)"
+  proof (rule UN_cases)
+    let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a)
+      (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}"
+    have "?A \<subseteq> ?B"
+    proof
+      fix x
+      assume "x \<in> ?A"
+      then obtain c d
+        where x:  "x = cbox c d"
+                  "\<And>i. i \<in> Basis \<Longrightarrow>
+                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+                        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast
+      show "x \<in> ?B"
+        unfolding image_iff x
+        apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
+        apply (rule arg_cong2 [where f = cbox])
+        using x(2) ab
+        apply (auto simp add: euclidean_eq_iff[where 'a='a])
+        by fastforce
+    qed
+    then show "finite ?A"
+      by (rule finite_subset) auto
+  next
+    fix s
+    assume "s \<in> ?A"
+    then obtain c d
+      where s: "s = cbox c d"
+               "\<And>i. i \<in> Basis \<Longrightarrow>
+                     c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+                     c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
+      by blast
+    show "P s"
+      unfolding s
+      apply (rule as[rule_format])
+      using ab s(2) by force
+    show "\<exists>a b. s = cbox a b"
+      unfolding s by auto
+    fix t
+    assume "t \<in> ?A"
+    then obtain e f where t:
+      "t = cbox e f"
+      "\<And>i. i \<in> Basis \<Longrightarrow>
+        e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+        e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i"
+      by blast
+    assume "s \<noteq> t"
+    then have "\<not> (c = e \<and> d = f)"
+      unfolding s t by auto
+    then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
+      unfolding euclidean_eq_iff[where 'a='a] by auto
+    then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
+      using s(2) t(2) apply fastforce
+      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
+    have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
+      by auto
+    show "interior s \<inter> interior t = {}"
+      unfolding s t interior_cbox
+    proof (rule *)
+      fix x
+      assume "x \<in> box c d" "x \<in> box e f"
+      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
+        unfolding mem_box using i'
+        by force+
+      show False  using s(2)[OF i']
+      proof safe
+        assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
+        show False
+          using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps)
+      next
+        assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
+        show False
+          using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
+      qed
+    qed
+  qed
+  also have "\<Union>?A = cbox a b"
+  proof (rule set_eqI,rule)
+    fix x
+    assume "x \<in> \<Union>?A"
+    then obtain c d where x:
+      "x \<in> cbox c d"
+      "\<And>i. i \<in> Basis \<Longrightarrow>
+        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
+        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
+      by blast
+    show "x\<in>cbox a b"
+      unfolding mem_box
+    proof safe
+      fix i :: 'a
+      assume i: "i \<in> Basis"
+      then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
+        using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto
+    qed
+  next
+    fix x
+    assume x: "x \<in> cbox a b"
+    have "\<forall>i\<in>Basis.
+      \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
+      (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d")
+      unfolding mem_box
+    proof
+      fix i :: 'a
+      assume i: "i \<in> Basis"
+      have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
+        using x[unfolded mem_box,THEN bspec, OF i] by auto
+      then show "\<exists>c d. ?P i c d"
+        by blast
+    qed
+    then show "x\<in>\<Union>?A"
+      unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff
+      apply auto
+      apply (rule_tac x="cbox xa xaa" in exI)
+      unfolding mem_box
+      apply auto
+      done
+  qed
+  finally show False
+    using assms by auto
+qed
+
+lemma interval_bisection:
+  fixes type :: "'a::euclidean_space"
+  assumes "P {}"
+    and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))"
+    and "\<not> P (cbox a (b::'a))"
+  obtains x where "x \<in> cbox a b"
+    and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
+proof -
+  have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
+    (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
+       2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
+  proof
+    show "?P x" for x
+    proof (cases "P (cbox (fst x) (snd x))")
+      case True
+      then show ?thesis by auto
+    next
+      case as: False
+      obtain c d where "\<not> P (cbox c d)"
+        "\<forall>i\<in>Basis.
+           fst x \<bullet> i \<le> c \<bullet> i \<and>
+           c \<bullet> i \<le> d \<bullet> i \<and>
+           d \<bullet> i \<le> snd x \<bullet> i \<and>
+           2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i"
+        by (rule interval_bisection_step[of P, OF assms(1-2) as])
+      then show ?thesis
+        apply -
+        apply (rule_tac x="(c,d)" in exI)
+        apply auto
+        done
+    qed
+  qed
+  then obtain f where f:
+    "\<forall>x.
+      \<not> P (cbox (fst x) (snd x)) \<longrightarrow>
+      \<not> P (cbox (fst (f x)) (snd (f x))) \<and>
+        (\<forall>i\<in>Basis.
+            fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and>
+            fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and>
+            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
+            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)"
+    apply -
+    apply (drule choice)
+    apply blast
+    done
+  define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
+  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
+    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
+    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
+  proof -
+    show "A 0 = a" "B 0 = b"
+      unfolding ab_def by auto
+    note S = ab_def funpow.simps o_def id_apply
+    show "?P n" for n
+    proof (induct n)
+      case 0
+      then show ?case
+        unfolding S
+        apply (rule f[rule_format]) using assms(3)
+        apply auto
+        done
+    next
+      case (Suc n)
+      show ?case
+        unfolding S
+        apply (rule f[rule_format])
+        using Suc
+        unfolding S
+        apply auto
+        done
+    qed
+  qed
+  note AB = this(1-2) conjunctD2[OF this(3),rule_format]
+
+  have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e"
+    if e: "0 < e" for e
+  proof -
+    obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
+      using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
+    show ?thesis
+    proof (rule exI [where x=n], clarify)
+      fix x y
+      assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
+      have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
+        unfolding dist_norm by(rule norm_le_l1)
+      also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
+      proof (rule setsum_mono)
+        fix i :: 'a
+        assume i: "i \<in> Basis"
+        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
+          using xy[unfolded mem_box,THEN bspec, OF i]
+          by (auto simp: inner_diff_left)
+      qed
+      also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
+        unfolding setsum_divide_distrib
+      proof (rule setsum_mono)
+        show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
+        proof (induct n)
+          case 0
+          then show ?case
+            unfolding AB by auto
+        next
+          case (Suc n)
+          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
+            using AB(4)[of i n] using i by auto
+          also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n"
+            using Suc by (auto simp add: field_simps)
+          finally show ?case .
+        qed
+      qed
+      also have "\<dots> < e"
+        using n using e by (auto simp add: field_simps)
+      finally show "dist x y < e" .
+    qed
+  qed
+  {
+    fix n m :: nat
+    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
+    proof (induction rule: inc_induct)
+      case (step i)
+      show ?case
+        using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto
+    qed simp
+  } note ABsubset = this
+  have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)"
+    by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv])
+      (metis nat.exhaust AB(1-3) assms(1,3))
+  then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
+    by blast
+  show thesis
+  proof (rule that[rule_format, of x0])
+    show "x0\<in>cbox a b"
+      using x0[of 0] unfolding AB .
+    fix e :: real
+    assume "e > 0"
+    from interv[OF this] obtain n
+      where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" ..
+    have "\<not> P (cbox (A n) (B n))"
+      apply (cases "0 < n")
+      using AB(3)[of "n - 1"] assms(3) AB(1-2)
+      apply auto
+      done
+    moreover have "cbox (A n) (B n) \<subseteq> ball x0 e"
+      using n using x0[of n] by auto
+    moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
+      unfolding AB(1-2)[symmetric] by (rule ABsubset) auto
+    ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
+      apply (rule_tac x="A n" in exI)
+      apply (rule_tac x="B n" in exI)
+      apply (auto simp: x0)
+      done
+  qed
+qed
+
+
+subsection \<open>Cousin's lemma.\<close>
+
+lemma fine_division_exists:
+  fixes a b :: "'a::euclidean_space"
+  assumes "gauge g"
+  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
+proof -
+  presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False"
+  then obtain p where "p tagged_division_of (cbox a b)" "g fine p"
+    by blast
+  then show thesis ..
+next
+  assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
+  obtain x where x:
+      "x \<in> (cbox a b)"
+      "\<And>e. 0 < e \<Longrightarrow>
+        \<exists>c d.
+          x \<in> cbox c d \<and>
+          cbox c d \<subseteq> ball x e \<and>
+          cbox c d \<subseteq> (cbox a b) \<and>
+          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
+    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as])
+    apply (simp add: fine_def)
+    apply (metis tagged_division_union fine_union)
+    apply (auto simp: )
+    done
+  obtain e where e: "e > 0" "ball x e \<subseteq> g x"
+    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
+  from x(2)[OF e(1)]
+  obtain c d where c_d: "x \<in> cbox c d"
+                        "cbox c d \<subseteq> ball x e"
+                        "cbox c d \<subseteq> cbox a b"
+                        "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
+    by blast
+  have "g fine {(x, cbox c d)}"
+    unfolding fine_def using e using c_d(2) by auto
+  then show False
+    using tagged_division_of_self[OF c_d(1)] using c_d by auto
+qed
+
+lemma fine_division_exists_real:
+  fixes a b :: real
+  assumes "gauge g"
+  obtains p where "p tagged_division_of {a .. b}" "g fine p"
+  by (metis assms box_real(2) fine_division_exists)
+
+subsection \<open>Basic theorems about integrals.\<close>
+
+lemma has_integral_unique:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral k1) i"
+    and "(f has_integral k2) i"
+  shows "k1 = k2"
+proof (rule ccontr)
+  let ?e = "norm (k1 - k2) / 2"
+  assume as: "k1 \<noteq> k2"
+  then have e: "?e > 0"
+    by auto
+  have lem: False
+    if f_k1: "(f has_integral k1) (cbox a b)"
+    and f_k2: "(f has_integral k2) (cbox a b)"
+    and "k1 \<noteq> k2"
+    for f :: "'n \<Rightarrow> 'a" and a b k1 k2
+  proof -
+    let ?e = "norm (k1 - k2) / 2"
+    from \<open>k1 \<noteq> k2\<close> have e: "?e > 0" by auto
+    obtain d1 where d1:
+        "gauge d1"
+        "\<And>p. p tagged_division_of cbox a b \<Longrightarrow>
+          d1 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2"
+      by (rule has_integralD[OF f_k1 e]) blast
+    obtain d2 where d2:
+        "gauge d2"
+        "\<And>p. p tagged_division_of cbox a b \<Longrightarrow>
+          d2 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2"
+      by (rule has_integralD[OF f_k2 e]) blast
+    obtain p where p:
+        "p tagged_division_of cbox a b"
+        "(\<lambda>x. d1 x \<inter> d2 x) fine p"
+      by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]])
+    let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
+    have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
+      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"]
+      by (auto simp add:algebra_simps norm_minus_commute)
+    also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
+      apply (rule add_strict_mono)
+      apply (rule_tac[!] d2(2) d1(2))
+      using p unfolding fine_def
+      apply auto
+      done
+    finally show False by auto
+  qed
+  {
+    presume "\<not> (\<exists>a b. i = cbox a b) \<Longrightarrow> False"
+    then show False
+      using as assms lem by blast
+  }
+  assume as: "\<not> (\<exists>a b. i = cbox a b)"
+  obtain B1 where B1:
+      "0 < B1"
+      "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow>
+        \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and>
+          norm (z - k1) < norm (k1 - k2) / 2"
+    by (rule has_integral_altD[OF assms(1) as,OF e]) blast
+  obtain B2 where B2:
+      "0 < B2"
+      "\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow>
+        \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and>
+          norm (z - k2) < norm (k1 - k2) / 2"
+    by (rule has_integral_altD[OF assms(2) as,OF e]) blast
+  have "\<exists>a b::'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> cbox a b"
+    apply (rule bounded_subset_cbox)
+    using bounded_Un bounded_ball
+    apply auto
+    done
+  then obtain a b :: 'n where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b"
+    by blast
+  obtain w where w:
+    "((\<lambda>x. if x \<in> i then f x else 0) has_integral w) (cbox a b)"
+    "norm (w - k1) < norm (k1 - k2) / 2"
+    using B1(2)[OF ab(1)] by blast
+  obtain z where z:
+    "((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b)"
+    "norm (z - k2) < norm (k1 - k2) / 2"
+    using B2(2)[OF ab(2)] by blast
+  have "z = w"
+    using lem[OF w(1) z(1)] by auto
+  then have "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)"
+    using norm_triangle_ineq4 [of "k1 - w" "k2 - z"]
+    by (auto simp add: norm_minus_commute)
+  also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
+    apply (rule add_strict_mono)
+    apply (rule_tac[!] z(2) w(2))
+    done
+  finally show False by auto
+qed
+
+lemma integral_unique [intro]: "(f has_integral y) k \<Longrightarrow> integral k f = y"
+  unfolding integral_def
+  by (rule some_equality) (auto intro: has_integral_unique)
+
+lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> ~ f integrable_on k \<and> y=0"
+  unfolding integral_def integrable_on_def
+  apply (erule subst)
+  apply (rule someI_ex)
+  by blast
+
+lemma has_integral_is_0:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "\<forall>x\<in>s. f x = 0"
+  shows "(f has_integral 0) s"
+proof -
+  have lem: "\<And>a b. \<And>f::'n \<Rightarrow> 'a.
+    (\<forall>x\<in>cbox a b. f(x) = 0) \<Longrightarrow> (f has_integral 0) (cbox a b)"
+    unfolding has_integral
+  proof clarify
+    fix a b e
+    fix f :: "'n \<Rightarrow> 'a"
+    assume as: "\<forall>x\<in>cbox a b. f x = 0" "0 < (e::real)"
+    have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
+      if p: "p tagged_division_of cbox a b" for p
+    proof -
+      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0"
+      proof (rule setsum.neutral, rule)
+        fix x
+        assume x: "x \<in> p"
+        have "f (fst x) = 0"
+          using tagged_division_ofD(2-3)[OF p, of "fst x" "snd x"] using as x by auto
+        then show "(\<lambda>(x, k). content k *\<^sub>R f x) x = 0"
+          apply (subst surjective_pairing[of x])
+          unfolding split_conv
+          apply auto
+          done
+      qed
+      then show ?thesis
+        using as by auto
+    qed
+    then show "\<exists>d. gauge d \<and>
+        (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)"
+      by auto
+  qed
+  {
+    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
+    with assms lem show ?thesis
+      by blast
+  }
+  have *: "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. 0)"
+    apply (rule ext)
+    using assms
+    apply auto
+    done
+  assume "\<not> (\<exists>a b. s = cbox a b)"
+  then show ?thesis
+    using lem
+    by (subst has_integral_alt) (force simp add: *)
+qed
+
+lemma has_integral_0[simp]: "((\<lambda>x::'n::euclidean_space. 0) has_integral 0) s"
+  by (rule has_integral_is_0) auto
+
+lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0"
+  using has_integral_unique[OF has_integral_0] by auto
+
+lemma has_integral_linear:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral y) s"
+    and "bounded_linear h"
+  shows "((h \<circ> f) has_integral ((h y))) s"
+proof -
+  interpret bounded_linear h
+    using assms(2) .
+  from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
+    by blast
+  have lem: "\<And>(f :: 'n \<Rightarrow> 'a) y a b.
+    (f has_integral y) (cbox a b) \<Longrightarrow> ((h \<circ> f) has_integral h y) (cbox a b)"
+    unfolding has_integral
+  proof (clarify, goal_cases)
+    case prems: (1 f y a b e)
+    from pos_bounded
+    obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
+      by blast
+    have "e / B > 0" using prems(2) B by simp
+    then obtain g
+      where g: "gauge g"
+               "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> g fine p \<Longrightarrow>
+                    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e / B"
+        using prems(1) by auto
+    {
+      fix p
+      assume as: "p tagged_division_of (cbox a b)" "g fine p"
+      have hc: "\<And>x k. h ((\<lambda>(x, k). content k *\<^sub>R f x) x) = (\<lambda>(x, k). h (content k *\<^sub>R f x)) x"
+        by auto
+      then have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = setsum (h \<circ> (\<lambda>(x, k). content k *\<^sub>R f x)) p"
+        unfolding o_def unfolding scaleR[symmetric] hc by simp
+      also have "\<dots> = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
+        using setsum[of "\<lambda>(x,k). content k *\<^sub>R f x" p] using as by auto
+      finally have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" .
+      then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) - h y) < e"
+        apply (simp add: diff[symmetric])
+        apply (rule le_less_trans[OF B(2)])
+        using g(2)[OF as] B(1)
+        apply (auto simp add: field_simps)
+        done
+    }
+    with g show ?case
+      by (rule_tac x=g in exI) auto
+  qed
+  {
+    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
+    then show ?thesis
+      using assms(1) lem by blast
+  }
+  assume as: "\<not> (\<exists>a b. s = cbox a b)"
+  then show ?thesis
+  proof (subst has_integral_alt, clarsimp)
+    fix e :: real
+    assume e: "e > 0"
+    have *: "0 < e/B" using e B(1) by simp
+    obtain M where M:
+      "M > 0"
+      "\<And>a b. ball 0 M \<subseteq> cbox a b \<Longrightarrow>
+        \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e / B"
+      using has_integral_altD[OF assms(1) as *] by blast
+    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) has_integral z) (cbox a b) \<and> norm (z - h y) < e)"
+    proof (rule_tac x=M in exI, clarsimp simp add: M, goal_cases)
+      case prems: (1 a b)
+      obtain z where z:
+        "((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b)"
+        "norm (z - y) < e / B"
+        using M(2)[OF prems(1)] by blast
+      have *: "(\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) = h \<circ> (\<lambda>x. if x \<in> s then f x else 0)"
+        using zero by auto
+      show ?case
+        apply (rule_tac x="h z" in exI)
+        apply (simp add: * lem z(1))
+        apply (metis B diff le_less_trans pos_less_divide_eq z(2))
+        done
+    qed
+  qed
+qed
+
+lemma has_integral_scaleR_left:
+  "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s"
+  using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def)
+
+lemma has_integral_mult_left:
+  fixes c :: "_ :: real_normed_algebra"
+  shows "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) s"
+  using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def)
+
+text\<open>The case analysis eliminates the condition @{term "f integrable_on s"} at the cost
+     of the type class constraint \<open>division_ring\<close>\<close>
+corollary integral_mult_left [simp]:
+  fixes c:: "'a::{real_normed_algebra,division_ring}"
+  shows "integral s (\<lambda>x. f x * c) = integral s f * c"
+proof (cases "f integrable_on s \<or> c = 0")
+  case True then show ?thesis
+    by (force intro: has_integral_mult_left)
+next
+  case False then have "~ (\<lambda>x. f x * c) integrable_on s"
+    using has_integral_mult_left [of "(\<lambda>x. f x * c)" _ s "inverse c"]
+    by (force simp add: mult.assoc)
+  with False show ?thesis by (simp add: not_integrable_integral)
+qed
+
+corollary integral_mult_right [simp]:
+  fixes c:: "'a::{real_normed_field}"
+  shows "integral s (\<lambda>x. c * f x) = c * integral s f"
+by (simp add: mult.commute [of c])
+
+corollary integral_divide [simp]:
+  fixes z :: "'a::real_normed_field"
+  shows "integral S (\<lambda>x. f x / z) = integral S (\<lambda>x. f x) / z"
+using integral_mult_left [of S f "inverse z"]
+  by (simp add: divide_inverse_commute)
+
+lemma has_integral_mult_right:
+  fixes c :: "'a :: real_normed_algebra"
+  shows "(f has_integral y) i \<Longrightarrow> ((\<lambda>x. c * f x) has_integral (c * y)) i"
+  using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def)
+
+lemma has_integral_cmul: "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s"
+  unfolding o_def[symmetric]
+  by (metis has_integral_linear bounded_linear_scaleR_right)
+
+lemma has_integral_cmult_real:
+  fixes c :: real
+  assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A"
+  shows "((\<lambda>x. c * f x) has_integral c * x) A"
+proof (cases "c = 0")
+  case True
+  then show ?thesis by simp
+next
+  case False
+  from has_integral_cmul[OF assms[OF this], of c] show ?thesis
+    unfolding real_scaleR_def .
+qed
+
+lemma has_integral_neg: "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral -k) s"
+  by (drule_tac c="-1" in has_integral_cmul) auto
+
+lemma has_integral_add:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "(f has_integral k) s"
+    and "(g has_integral l) s"
+  shows "((\<lambda>x. f x + g x) has_integral (k + l)) s"
+proof -
+  have lem: "((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)"
+    if f_k: "(f has_integral k) (cbox a b)"
+    and g_l: "(g has_integral l) (cbox a b)"
+    for f :: "'n \<Rightarrow> 'a" and g a b k l
+    unfolding has_integral
+  proof clarify
+    fix e :: real
+    assume e: "e > 0"
+    then have *: "e / 2 > 0"
+      by auto
+    obtain d1 where d1:
+      "gauge d1"
+      "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d1 fine p \<Longrightarrow>
+        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) < e / 2"
+      using has_integralD[OF f_k *] by blast
+    obtain d2 where d2:
+      "gauge d2"
+      "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d2 fine p \<Longrightarrow>
+        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l) < e / 2"
+      using has_integralD[OF g_l *] by blast
+    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+              norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)"
+    proof (rule exI [where x="\<lambda>x. (d1 x) \<inter> (d2 x)"], clarsimp simp add: gauge_inter[OF d1(1) d2(1)])
+      fix p
+      assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d1 x \<inter> d2 x) fine p"
+      have *: "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) =
+        (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)"
+        unfolding scaleR_right_distrib setsum.distrib[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,symmetric]
+        by (rule setsum.cong) auto
+      from as have fine: "d1 fine p" "d2 fine p"
+        unfolding fine_inter by auto
+      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) =
+            norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
+        unfolding * by (auto simp add: algebra_simps)
+      also have "\<dots> < e/2 + e/2"
+        apply (rule le_less_trans[OF norm_triangle_ineq])
+        using as d1 d2 fine
+        apply (blast intro: add_strict_mono)
+        done
+      finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e"
+        by auto
+    qed
+  qed
+  {
+    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
+    then show ?thesis
+      using assms lem by force
+  }
+  assume as: "\<not> (\<exists>a b. s = cbox a b)"
+  then show ?thesis
+  proof (subst has_integral_alt, clarsimp, goal_cases)
+    case (1 e)
+    then have *: "e / 2 > 0"
+      by auto
+    from has_integral_altD[OF assms(1) as *]
+    obtain B1 where B1:
+        "0 < B1"
+        "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow>
+          \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - k) < e / 2"
+      by blast
+    from has_integral_altD[OF assms(2) as *]
+    obtain B2 where B2:
+        "0 < B2"
+        "\<And>a b. ball 0 B2 \<subseteq> (cbox a b) \<Longrightarrow>
+          \<exists>z. ((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b) \<and> norm (z - l) < e / 2"
+      by blast
+    show ?case
+    proof (rule_tac x="max B1 B2" in exI, clarsimp simp add: max.strict_coboundedI1 B1)
+      fix a b
+      assume "ball 0 (max B1 B2) \<subseteq> cbox a (b::'n)"
+      then have *: "ball 0 B1 \<subseteq> cbox a (b::'n)" "ball 0 B2 \<subseteq> cbox a (b::'n)"
+        by auto
+      obtain w where w:
+        "((\<lambda>x. if x \<in> s then f x else 0) has_integral w) (cbox a b)"
+        "norm (w - k) < e / 2"
+        using B1(2)[OF *(1)] by blast
+      obtain z where z:
+        "((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b)"
+        "norm (z - l) < e / 2"
+        using B2(2)[OF *(2)] by blast
+      have *: "\<And>x. (if x \<in> s then f x + g x else 0) =
+        (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)"
+        by auto
+      show "\<exists>z. ((\<lambda>x. if x \<in> s then f x + g x else 0) has_integral z) (cbox a b) \<and> norm (z - (k + l)) < e"
+        apply (rule_tac x="w + z" in exI)
+        apply (simp add: lem[OF w(1) z(1), unfolded *[symmetric]])
+        using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2)
+        apply (auto simp add: field_simps)
+        done
+    qed
+  qed
+qed
+
+lemma has_integral_sub:
+  "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow>
+    ((\<lambda>x. f x - g x) has_integral (k - l)) s"
+  using has_integral_add[OF _ has_integral_neg, of f k s g l]
+  by (auto simp: algebra_simps)
+
+lemma integral_0 [simp]:
+  "integral s (\<lambda>x::'n::euclidean_space. 0::'m::real_normed_vector) = 0"
+  by (rule integral_unique has_integral_0)+
+
+lemma integral_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
+    integral s (\<lambda>x. f x + g x) = integral s f + integral s g"
+  by (rule integral_unique) (metis integrable_integral has_integral_add)
+
+lemma integral_cmul [simp]: "integral s (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral s f"
+proof (cases "f integrable_on s \<or> c = 0")
+  case True with has_integral_cmul show ?thesis by force
+next
+  case False then have "~ (\<lambda>x. c *\<^sub>R f x) integrable_on s"
+    using has_integral_cmul [of "(\<lambda>x. c *\<^sub>R f x)" _ s "inverse c"]
+    by force
+  with False show ?thesis by (simp add: not_integrable_integral)
+qed
+
+lemma integral_neg [simp]: "integral s (\<lambda>x. - f x) = - integral s f"
+proof (cases "f integrable_on s")
+  case True then show ?thesis
+    by (simp add: has_integral_neg integrable_integral integral_unique)
+next
+  case False then have "~ (\<lambda>x. - f x) integrable_on s"
+    using has_integral_neg [of "(\<lambda>x. - f x)" _ s ]
+    by force
+  with False show ?thesis by (simp add: not_integrable_integral)
+qed
+
+lemma integral_diff: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
+    integral s (\<lambda>x. f x - g x) = integral s f - integral s g"
+  by (rule integral_unique) (metis integrable_integral has_integral_sub)
+
+lemma integrable_0: "(\<lambda>x. 0) integrable_on s"
+  unfolding integrable_on_def using has_integral_0 by auto
+
+lemma integrable_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_add)
+
+lemma integrable_cmul: "f integrable_on s \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_cmul)
+
+lemma integrable_on_cmult_iff:
+  fixes c :: real
+  assumes "c \<noteq> 0"
+  shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
+  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] \<open>c \<noteq> 0\<close>
+  by auto
+
+lemma integrable_on_cmult_left:
+  assumes "f integrable_on s"
+  shows "(\<lambda>x. of_real c * f x) integrable_on s"
+    using integrable_cmul[of f s "of_real c"] assms
+    by (simp add: scaleR_conv_of_real)
+
+lemma integrable_neg: "f integrable_on s \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_neg)
+
+lemma integrable_diff:
+  "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_sub)
+
+lemma integrable_linear:
+  "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) integrable_on s"
+  unfolding integrable_on_def by(auto intro: has_integral_linear)
+
+lemma integral_linear:
+  "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> integral s (h \<circ> f) = h (integral s f)"
+  apply (rule has_integral_unique [where i=s and f = "h \<circ> f"])
+  apply (simp_all add: integrable_integral integrable_linear has_integral_linear )
+  done
+
+lemma integral_component_eq[simp]:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "f integrable_on s"
+  shows "integral s (\<lambda>x. f x \<bullet> k) = integral s f \<bullet> k"
+  unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] ..
+
+lemma has_integral_setsum:
+  assumes "finite t"
+    and "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
+  shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum i t)) s"
+  using assms(1) subset_refl[of t]
+proof (induct rule: finite_subset_induct)
+  case empty
+  then show ?case by auto
+next
+  case (insert x F)
+  with assms show ?case
+    by (simp add: has_integral_add)
+qed
+
+lemma integral_setsum:
+  "\<lbrakk>finite t;  \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow>
+   integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t"
+  by (auto intro: has_integral_setsum integrable_integral)
+
+lemma integrable_setsum:
+  "\<lbrakk>finite t;  \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s"
+  unfolding integrable_on_def
+  apply (drule bchoice)
+  using has_integral_setsum[of t]
+  apply auto
+  done
+
+lemma has_integral_eq:
+  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
+    and "(f has_integral k) s"
+  shows "(g has_integral k) s"
+  using has_integral_sub[OF assms(2), of "\<lambda>x. f x - g x" 0]
+  using has_integral_is_0[of s "\<lambda>x. f x - g x"]
+  using assms(1)
+  by auto
+
+lemma integrable_eq: "(\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f integrable_on s \<Longrightarrow> g integrable_on s"
+  unfolding integrable_on_def
+  using has_integral_eq[of s f g] has_integral_eq by blast
+
+lemma has_integral_cong:
+  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
+  shows "(f has_integral i) s = (g has_integral i) s"
+  using has_integral_eq[of s f g] has_integral_eq[of s g f] assms
+  by auto
+
+lemma integral_cong:
+  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
+  shows "integral s f = integral s g"
+  unfolding integral_def
+by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq)
+
+lemma integrable_on_cmult_left_iff [simp]:
+  assumes "c \<noteq> 0"
+  shows "(\<lambda>x. of_real c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
+        (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then have "(\<lambda>x. of_real (1 / c) * (of_real c * f x)) integrable_on s"
+    using integrable_cmul[of "\<lambda>x. of_real c * f x" s "1 / of_real c"]
+    by (simp add: scaleR_conv_of_real)
+  then have "(\<lambda>x. (of_real (1 / c) * of_real c * f x)) integrable_on s"
+    by (simp add: algebra_simps)
+  with \<open>c \<noteq> 0\<close> show ?rhs
+    by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult)
+qed (blast intro: integrable_on_cmult_left)
+
+lemma integrable_on_cmult_right:
+  fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
+  assumes "f integrable_on s"
+  shows "(\<lambda>x. f x * of_real c) integrable_on s"
+using integrable_on_cmult_left [OF assms] by (simp add: mult.commute)
+
+lemma integrable_on_cmult_right_iff [simp]:
+  fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
+  assumes "c \<noteq> 0"
+  shows "(\<lambda>x. f x * of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
+using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute)
+
+lemma integrable_on_cdivide:
+  fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
+  assumes "f integrable_on s"
+  shows "(\<lambda>x. f x / of_real c) integrable_on s"
+by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
+
+lemma integrable_on_cdivide_iff [simp]:
+  fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
+  assumes "c \<noteq> 0"
+  shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
+by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
+
+lemma has_integral_null [intro]:
+  assumes "content(cbox a b) = 0"
+  shows "(f has_integral 0) (cbox a b)"
+proof -
+  have "gauge (\<lambda>x. ball x 1)"
+    by auto
+  moreover
+  {
+    fix e :: real
+    fix p
+    assume e: "e > 0"
+    assume p: "p tagged_division_of (cbox a b)"
+    have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) = 0"
+      unfolding norm_eq_zero diff_0_right
+      using setsum_content_null[OF assms(1) p, of f] .
+    then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
+      using e by auto
+  }
+  ultimately show ?thesis
+    by (auto simp: has_integral)
+qed
+
+lemma has_integral_null_real [intro]:
+  assumes "content {a .. b::real} = 0"
+  shows "(f has_integral 0) {a .. b}"
+  by (metis assms box_real(2) has_integral_null)
+
+lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \<Longrightarrow> (f has_integral i) (cbox a b) \<longleftrightarrow> i = 0"
+  by (auto simp add: has_integral_null dest!: integral_unique)
+
+lemma integral_null [simp]: "content (cbox a b) = 0 \<Longrightarrow> integral (cbox a b) f = 0"
+  by (metis has_integral_null integral_unique)
+
+lemma integrable_on_null [intro]: "content (cbox a b) = 0 \<Longrightarrow> f integrable_on (cbox a b)"
+  by (simp add: has_integral_integrable)
+
+lemma has_integral_empty[intro]: "(f has_integral 0) {}"
+  by (simp add: has_integral_is_0)
+
+lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \<longleftrightarrow> i = 0"
+  by (auto simp add: has_integral_empty has_integral_unique)
+
+lemma integrable_on_empty[intro]: "f integrable_on {}"
+  unfolding integrable_on_def by auto
+
+lemma integral_empty[simp]: "integral {} f = 0"
+  by (rule integral_unique) (rule has_integral_empty)
+
+lemma has_integral_refl[intro]:
+  fixes a :: "'a::euclidean_space"
+  shows "(f has_integral 0) (cbox a a)"
+    and "(f has_integral 0) {a}"
+proof -
+  have *: "{a} = cbox a a"
+    apply (rule set_eqI)
+    unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a]
+    apply safe
+    prefer 3
+    apply (erule_tac x=b in ballE)
+    apply (auto simp add: field_simps)
+    done
+  show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}"
+    unfolding *
+    apply (rule_tac[!] has_integral_null)
+    unfolding content_eq_0_interior
+    unfolding interior_cbox
+    using box_sing
+    apply auto
+    done
+qed
+
+lemma integrable_on_refl[intro]: "f integrable_on cbox a a"
+  unfolding integrable_on_def by auto
+
+lemma integral_refl [simp]: "integral (cbox a a) f = 0"
+  by (rule integral_unique) auto
+
+lemma integral_singleton [simp]: "integral {a} f = 0"
+  by auto
+
+lemma integral_blinfun_apply:
+  assumes "f integrable_on s"
+  shows "integral s (\<lambda>x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)"
+  by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def)
+
+lemma blinfun_apply_integral:
+  assumes "f integrable_on s"
+  shows "blinfun_apply (integral s f) x = integral s (\<lambda>y. blinfun_apply (f y) x)"
+  by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong)
+
+lemma has_integral_componentwise_iff:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
+  shows "(f has_integral y) A \<longleftrightarrow> (\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
+proof safe
+  fix b :: 'b assume "(f has_integral y) A"
+  from has_integral_linear[OF this(1) bounded_linear_component, of b]
+    show "((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A" by (simp add: o_def)
+next
+  assume "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
+  hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral ((y \<bullet> b) *\<^sub>R b)) A"
+    by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
+  hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. (y \<bullet> b) *\<^sub>R b)) A"
+    by (intro has_integral_setsum) (simp_all add: o_def)
+  thus "(f has_integral y) A" by (simp add: euclidean_representation)
+qed
+
+lemma has_integral_componentwise:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
+  shows "(\<And>b. b \<in> Basis \<Longrightarrow> ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A) \<Longrightarrow> (f has_integral y) A"
+  by (subst has_integral_componentwise_iff) blast
+
+lemma integrable_componentwise_iff:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
+  shows "f integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)"
+proof
+  assume "f integrable_on A"
+  then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def)
+  hence "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
+    by (subst (asm) has_integral_componentwise_iff)
+  thus "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" by (auto simp: integrable_on_def)
+next
+  assume "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)"
+  then obtain y where "\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral y b) A"
+    unfolding integrable_on_def by (subst (asm) bchoice_iff) blast
+  hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral (y b *\<^sub>R b)) A"
+    by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
+  hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. y b *\<^sub>R b)) A"
+    by (intro has_integral_setsum) (simp_all add: o_def)
+  thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation)
+qed
+
+lemma integrable_componentwise:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
+  shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) integrable_on A) \<Longrightarrow> f integrable_on A"
+  by (subst integrable_componentwise_iff) blast
+
+lemma integral_componentwise:
+  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
+  assumes "f integrable_on A"
+  shows "integral A f = (\<Sum>b\<in>Basis. integral A (\<lambda>x. (f x \<bullet> b) *\<^sub>R b))"
+proof -
+  from assms have integrable: "\<forall>b\<in>Basis. (\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. (f x \<bullet> b)) integrable_on A"
+    by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI)
+       (simp_all add: bounded_linear_scaleR_left)
+  have "integral A f = integral A (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b)"
+    by (simp add: euclidean_representation)
+  also from integrable have "\<dots> = (\<Sum>a\<in>Basis. integral A (\<lambda>x. (f x \<bullet> a) *\<^sub>R a))"
+    by (subst integral_setsum) (simp_all add: o_def)
+  finally show ?thesis .
+qed
+
+lemma integrable_component:
+  "f integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (y :: 'b :: euclidean_space)) integrable_on A"
+  by (drule integrable_linear[OF _ bounded_linear_component[of y]]) (simp add: o_def)
+
+
+
+subsection \<open>Cauchy-type criterion for integrability.\<close>
+
+(* XXXXXXX *)
+lemma integrable_cauchy:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::{real_normed_vector,complete_space}"
+  shows "f integrable_on cbox a b \<longleftrightarrow>
+    (\<forall>e>0.\<exists>d. gauge d \<and>
+      (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<and> d fine p1 \<and>
+        p2 tagged_division_of (cbox a b) \<and> d fine p2 \<longrightarrow>
+        norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 -
+        setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) < e))"
+  (is "?l = (\<forall>e>0. \<exists>d. ?P e d)")
+proof
+  assume ?l
+  then guess y unfolding integrable_on_def has_integral .. note y=this
+  show "\<forall>e>0. \<exists>d. ?P e d"
+  proof (clarify, goal_cases)
+    case (1 e)
+    then have "e/2 > 0" by auto
+    then guess d
+      apply -
+      apply (drule y[rule_format])
+      apply (elim exE conjE)
+      done
+    note d=this[rule_format]
+    show ?case
+    proof (rule_tac x=d in exI, clarsimp simp: d)
+      fix p1 p2
+      assume as: "p1 tagged_division_of (cbox a b)" "d fine p1"
+                 "p2 tagged_division_of (cbox a b)" "d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+        apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm])
+        using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] .
+    qed
+  qed
+next
+  assume "\<forall>e>0. \<exists>d. ?P e d"
+  then have "\<forall>n::nat. \<exists>d. ?P (inverse(of_nat (n + 1))) d"
+    by auto
+  from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
+  have "\<And>n. gauge (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}})"
+    apply (rule gauge_inters)
+    using d(1)
+    apply auto
+    done
+  then have "\<forall>n. \<exists>p. p tagged_division_of (cbox a b) \<and> (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}}) fine p"
+    by (meson fine_division_exists)
+  from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]]
+  have dp: "\<And>i n. i\<le>n \<Longrightarrow> d i fine p n"
+    using p(2) unfolding fine_inters by auto
+  have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
+  proof (rule CauchyI, goal_cases)
+    case (1 e)
+    then guess N unfolding real_arch_inverse[of e] .. note N=this
+    show ?case
+      apply (rule_tac x=N in exI)
+    proof clarify
+      fix m n
+      assume mn: "N \<le> m" "N \<le> n"
+      have *: "N = (N - 1) + 1" using N by auto
+      show "norm ((\<Sum>(x, k)\<in>p m. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p n. content k *\<^sub>R f x)) < e"
+        apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]])
+        apply(subst *)
+        using dp p(1) mn d(2) by auto
+    qed
+  qed
+  then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D]
+  show ?l
+    unfolding integrable_on_def has_integral
+  proof (rule_tac x=y in exI, clarify)
+    fix e :: real
+    assume "e>0"
+    then have *:"e/2 > 0" by auto
+    then guess N1 unfolding real_arch_inverse[of "e/2"] .. note N1=this
+    then have N1': "N1 = N1 - 1 + 1"
+      by auto
+    guess N2 using y[OF *] .. note N2=this
+    have "gauge (d (N1 + N2))"
+      using d by auto
+    moreover
+    {
+      fix q
+      assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q"
+      have *: "inverse (of_nat (N1 + N2 + 1)) < e / 2"
+        apply (rule less_trans)
+        using N1
+        apply auto
+        done
+      have "norm ((\<Sum>(x, k)\<in>q. content k *\<^sub>R f x) - y) < e"
+        apply (rule norm_triangle_half_r)
+        apply (rule less_trans[OF _ *])
+        apply (subst N1', rule d(2)[of "p (N1+N2)"])
+        using N1' as(1) as(2) dp
+        apply (simp add: \<open>\<forall>x. p x tagged_division_of cbox a b \<and> (\<lambda>xa. \<Inter>{d i xa |i. i \<in> {0..x}}) fine p x\<close>)
+        using N2 le_add2 by blast
+    }
+    ultimately show "\<exists>d. gauge d \<and>
+      (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e)"
+      by (rule_tac x="d (N1 + N2)" in exI) auto
+  qed
+qed
+
+
+subsection \<open>Additivity of integral on abutting intervals.\<close>
+
+lemma tagged_division_split_left_inj:
+  fixes x1 :: "'a::euclidean_space"
+  assumes d: "d tagged_division_of i"
+    and k12: "(x1, k1) \<in> d"
+             "(x2, k2) \<in> d"
+             "k1 \<noteq> k2"
+             "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
+             "k \<in> Basis"
+  shows "content (k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
+proof -
+  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
+    by force
+  show ?thesis
+    using k12
+    by (fastforce intro!:  division_split_left_inj[OF division_of_tagged_division[OF d]] *)
+qed
+
+lemma tagged_division_split_right_inj:
+  fixes x1 :: "'a::euclidean_space"
+  assumes d: "d tagged_division_of i"
+    and k12: "(x1, k1) \<in> d"
+             "(x2, k2) \<in> d"
+             "k1 \<noteq> k2"
+             "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
+             "k \<in> Basis"
+  shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
+proof -
+  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
+    by force
+  show ?thesis
+    using k12
+    by (fastforce intro!:  division_split_right_inj[OF division_of_tagged_division[OF d]] *)
+qed
+
+lemma has_integral_split:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes fi: "(f has_integral i) (cbox a b \<inter> {x. x\<bullet>k \<le> c})"
+      and fj: "(f has_integral j) (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+      and k: "k \<in> Basis"
+  shows "(f has_integral (i + j)) (cbox a b)"
+proof (unfold has_integral, rule, rule, goal_cases)
+  case (1 e)
+  then have e: "e/2 > 0"
+    by auto
+    obtain d1
+    where d1: "gauge d1"
+      and d1norm:
+        "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c};
+               d1 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. content k *\<^sub>R f x) - i) < e / 2"
+       apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e])
+       apply (simp add: interval_split[symmetric] k)
+       done
+    obtain d2
+    where d2: "gauge d2"
+      and d2norm:
+        "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k};
+               d2 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. content k *\<^sub>R f x) - j) < e / 2"
+       apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e])
+       apply (simp add: interval_split[symmetric] k)
+       done
+  let ?d = "\<lambda>x. if x\<bullet>k = c then (d1 x \<inter> d2 x) else ball x \<bar>x\<bullet>k - c\<bar> \<inter> d1 x \<inter> d2 x"
+  have "gauge ?d"
+    using d1 d2 unfolding gauge_def by auto
+  then show ?case
+  proof (rule_tac x="?d" in exI, safe)
+    fix p
+    assume "p tagged_division_of (cbox a b)" "?d fine p"
+    note p = this tagged_division_ofD[OF this(1)]
+    have xk_le_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<le> c"
+    proof -
+      fix x kk
+      assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}"
+      show "x\<bullet>k \<le> c"
+      proof (rule ccontr)
+        assume **: "\<not> ?thesis"
+        from this[unfolded not_le]
+        have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
+          using p(2)[unfolded fine_def, rule_format,OF as] by auto
+        with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<le> c"
+          by blast
+        then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>"
+          using Basis_le_norm[OF k, of "x - y"]
+          by (auto simp add: dist_norm inner_diff_left intro: le_less_trans)
+        with y show False
+          using ** by (auto simp add: field_simps)
+      qed
+    qed
+    have xk_ge_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<ge> c"
+    proof -
+      fix x kk
+      assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}"
+      show "x\<bullet>k \<ge> c"
+      proof (rule ccontr)
+        assume **: "\<not> ?thesis"
+        from this[unfolded not_le] have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
+          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
+        with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<ge> c"
+          by blast
+        then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>"
+          using Basis_le_norm[OF k, of "x - y"]
+          by (auto simp add: dist_norm inner_diff_left intro: le_less_trans)
+        with y show False
+          using ** by (auto simp add: field_simps)
+      qed
+    qed
+
+    have lem1: "\<And>f P Q. (\<forall>x k. (x, k) \<in> {(x, f k) | x k. P x k} \<longrightarrow> Q x k) \<longleftrightarrow>
+                         (\<forall>x k. P x k \<longrightarrow> Q x (f k))"
+      by auto
+    have fin_finite: "finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}" if "finite s" for f s P
+    proof -
+      from that have "finite ((\<lambda>(x, k). (x, f k)) ` s)"
+        by auto
+      then show ?thesis
+        by (rule rev_finite_subset) auto
+    qed
+    { fix g :: "'a set \<Rightarrow> 'a set"
+      fix i :: "'a \<times> 'a set"
+      assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
+      then obtain x k where xk:
+              "i = (x, g k)"  "(x, k) \<in> p"
+              "(x, g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
+          by auto
+      have "content (g k) = 0"
+        using xk using content_empty by auto
+      then have "(\<lambda>(x, k). content k *\<^sub>R f x) i = 0"
+        unfolding xk split_conv by auto
+    } note [simp] = this
+    have lem3: "\<And>g :: 'a set \<Rightarrow> 'a set. finite p \<Longrightarrow>
+                  setsum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} =
+                  setsum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)"
+      by (rule setsum.mono_neutral_left) auto
+    let ?M1 = "{(x, kk \<inter> {x. x\<bullet>k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
+    have d1_fine: "d1 fine ?M1"
+      by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm)
+    have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2"
+    proof (rule d1norm [OF tagged_division_ofI d1_fine])
+      show "finite ?M1"
+        by (rule fin_finite p(3))+
+      show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = cbox a b \<inter> {x. x\<bullet>k \<le> c}"
+        unfolding p(8)[symmetric] by auto
+      fix x l
+      assume xl: "(x, l) \<in> ?M1"
+      then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this
+      show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<le> c}"
+        unfolding xl'
+        using p(4-6)[OF xl'(3)] using xl'(4)
+        using xk_le_c[OF xl'(3-4)] by auto
+      show "\<exists>a b. l = cbox a b"
+        unfolding xl'
+        using p(6)[OF xl'(3)]
+        by (fastforce simp add: interval_split[OF k,where c=c])
+      fix y r
+      let ?goal = "interior l \<inter> interior r = {}"
+      assume yr: "(y, r) \<in> ?M1"
+      then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this
+      assume as: "(x, l) \<noteq> (y, r)"
+      show "interior l \<inter> interior r = {}"
+      proof (cases "l' = r' \<longrightarrow> x' = y'")
+        case False
+        then show ?thesis
+          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      next
+        case True
+        then have "l' \<noteq> r'"
+          using as unfolding xl' yr' by auto
+        then show ?thesis
+          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      qed
+    qed
+    moreover
+    let ?M2 = "{(x,kk \<inter> {x. x\<bullet>k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
+    have d2_fine: "d2 fine ?M2"
+      by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm)
+    have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2"
+    proof (rule d2norm [OF tagged_division_ofI d2_fine])
+      show "finite ?M2"
+        by (rule fin_finite p(3))+
+      show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = cbox a b \<inter> {x. x\<bullet>k \<ge> c}"
+        unfolding p(8)[symmetric] by auto
+      fix x l
+      assume xl: "(x, l) \<in> ?M2"
+      then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this
+      show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
+        unfolding xl'
+        using p(4-6)[OF xl'(3)] xl'(4) xk_ge_c[OF xl'(3-4)]
+        by auto
+      show "\<exists>a b. l = cbox a b"
+        unfolding xl'
+        using p(6)[OF xl'(3)]
+        by (fastforce simp add: interval_split[OF k, where c=c])
+      fix y r
+      let ?goal = "interior l \<inter> interior r = {}"
+      assume yr: "(y, r) \<in> ?M2"
+      then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this
+      assume as: "(x, l) \<noteq> (y, r)"
+      show "interior l \<inter> interior r = {}"
+      proof (cases "l' = r' \<longrightarrow> x' = y'")
+        case False
+        then show ?thesis
+          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      next
+        case True
+        then have "l' \<noteq> r'"
+          using as unfolding xl' yr' by auto
+        then show ?thesis
+          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
+      qed
+    qed
+    ultimately
+    have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2"
+      using norm_add_less by blast
+    also {
+      have eq0: "\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0"
+        using scaleR_zero_left by auto
+      have cont_eq: "\<And>g. (\<lambda>(x,l). content l *\<^sub>R f x) \<circ> (\<lambda>(x,l). (x,g l)) = (\<lambda>(x,l). content (g l) *\<^sub>R f x)"
+        by auto
+      have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) =
+        (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)"
+        by auto
+      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
+        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
+        unfolding lem3[OF p(3)]
+        by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)]
+              simp: cont_eq)+
+      also note setsum.distrib[symmetric]
+      also have "\<And>x. x \<in> p \<Longrightarrow>
+                    (\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x +
+                    (\<lambda>(x,ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x =
+                    (\<lambda>(x,ka). content ka *\<^sub>R f x) x"
+      proof clarify
+        fix a b
+        assume "(a, b) \<in> p"
+        from p(6)[OF this] guess u v by (elim exE) note uv=this
+        then show "content (b \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f a =
+          content b *\<^sub>R f a"
+          unfolding scaleR_left_distrib[symmetric]
+          unfolding uv content_split[OF k,of u v c]
+          by auto
+      qed
+      note setsum.cong [OF _ this]
+      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x \<bullet> k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
+        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x \<bullet> k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
+        (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)"
+        by auto
+    }
+    finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e"
+      by auto
+  qed
+qed
+
+
+subsection \<open>A sort of converse, integrability on subintervals.\<close>
+
+lemma tagged_division_union_interval:
+  fixes a :: "'a::euclidean_space"
+  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
+    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+    and k: "k \<in> Basis"
+  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
+proof -
+  have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
+    by auto
+  show ?thesis
+    apply (subst *)
+    apply (rule tagged_division_union[OF assms(1-2)])
+    unfolding interval_split[OF k] interior_cbox
+    using k
+    apply (auto simp add: box_def elim!: ballE[where x=k])
+    done
+qed
+
+lemma tagged_division_union_interval_real:
+  fixes a :: real
+  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
+    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
+    and k: "k \<in> Basis"
+  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
+  using assms
+  unfolding box_real[symmetric]
+  by (rule tagged_division_union_interval)
+
+lemma has_integral_separate_sides:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "(f has_integral i) (cbox a b)"
+    and "e > 0"
+    and k: "k \<in> Basis"
+  obtains d where "gauge d"
+    "\<forall>p1 p2. p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and>
+        p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2 \<longrightarrow>
+        norm ((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e"
+proof -
+  guess d using has_integralD[OF assms(1-2)] . note d=this
+  { fix p1 p2
+    assume "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1"
+    note p1=tagged_division_ofD[OF this(1)] this
+    assume "p2 tagged_division_of (cbox a b) \<inter> {x. c \<le> x \<bullet> k}" "d fine p2"
+    note p2=tagged_division_ofD[OF this(1)] this
+    note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
+    { fix a b
+      assume ab: "(a, b) \<in> p1 \<inter> p2"
+      have "(a, b) \<in> p1"
+        using ab by auto
+      with p1 obtain u v where uv: "b = cbox u v" by auto
+      have "b \<subseteq> {x. x\<bullet>k = c}"
+        using ab p1(3)[of a b] p2(3)[of a b] by fastforce
+      moreover
+      have "interior {x::'a. x \<bullet> k = c} = {}"
+      proof (rule ccontr)
+        assume "\<not> ?thesis"
+        then obtain x where x: "x \<in> interior {x::'a. x\<bullet>k = c}"
+          by auto
+        then guess e unfolding mem_interior .. note e=this
+        have x: "x\<bullet>k = c"
+          using x interior_subset by fastforce
+        have *: "\<And>i. i \<in> Basis \<Longrightarrow> \<bar>(x - (x + (e / 2) *\<^sub>R k)) \<bullet> i\<bar> = (if i = k then e/2 else 0)"
+          using e k by (auto simp: inner_simps inner_not_same_Basis)
+        have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
+              (\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))"
+          using "*" by (blast intro: setsum.cong)
+        also have "\<dots> < e"
+          apply (subst setsum.delta)
+          using e
+          apply auto
+          done
+        finally have "x + (e/2) *\<^sub>R k \<in> ball x e"
+          unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1])
+        then have "x + (e/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}"
+          using e by auto
+        then show False
+          unfolding mem_Collect_eq using e x k by (auto simp: inner_simps)
+      qed
+      ultimately have "content b = 0"
+        unfolding uv content_eq_0_interior
+        using interior_mono by blast
+      then have "content b *\<^sub>R f a = 0"
+        by auto
+    }
+    then have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) =
+               norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
+      by (subst setsum.union_inter_neutral) (auto simp: p1 p2)
+    also have "\<dots> < e"
+      by (rule k d(2) p12 fine_union p1 p2)+
+    finally have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" .
+   }
+  then show ?thesis
+    by (auto intro: that[of d] d elim: )
+qed
+
+lemma integrable_split[intro]:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
+  assumes "f integrable_on cbox a b"
+    and k: "k \<in> Basis"
+  shows "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<le> c})" (is ?t1)
+    and "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" (is ?t2)
+proof -
+  guess y using assms(1) unfolding integrable_on_def .. note y=this
+  define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)"
+  define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i)"
+  show ?t1 ?t2
+    unfolding interval_split[OF k] integrable_cauchy
+    unfolding interval_split[symmetric,OF k]
+  proof (rule_tac[!] allI impI)+
+    fix e :: real
+    assume "e > 0"
+    then have "e/2>0"
+      by auto
+    from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format]
+    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<inter> A \<and> d fine p1 \<and>
+      p2 tagged_division_of (cbox a b) \<inter> A \<and> d fine p2 \<longrightarrow>
+      norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
+    show "?P {x. x \<bullet> k \<le> c}"
+    proof (rule_tac x=d in exI, clarsimp simp add: d)
+      fix p1 p2
+      assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1"
+                 "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+      proof (rule fine_division_exists[OF d(1), of a' b] )
+        fix p
+        assume "p tagged_division_of cbox a' b" "d fine p"
+        then show ?thesis
+          using as norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
+          unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
+          by (auto simp add: algebra_simps)
+      qed
+    qed
+    show "?P {x. x \<bullet> k \<ge> c}"
+    proof (rule_tac x=d in exI, clarsimp simp add: d)
+      fix p1 p2
+      assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p1"
+                 "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p2"
+      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
+      proof (rule fine_division_exists[OF d(1), of a b'] )
+        fix p
+        assume "p tagged_division_of cbox a b'" "d fine p"
+        then show ?thesis
+          using as norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
+          unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
+          by (auto simp add: algebra_simps)
+      qed
+    qed
+  qed
+qed
+
+lemma operative_integral:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  shows "comm_monoid.operative (lift_option op +) (Some 0) (\<lambda>i. if f integrable_on i then Some (integral i f) else None)"
+  unfolding comm_monoid.operative_def[OF add.comm_monoid_lift_option]
+proof safe
+  fix a b c
+  fix k :: 'a
+  assume k: "k \<in> Basis"
+  show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) =
+        lift_option op + (if f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c} then Some (integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f) else None)
+        (if f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k} then Some (integral (cbox a b \<inter> {x. c \<le> x \<bullet> k}) f) else None)"
+  proof (cases "f integrable_on cbox a b")
+    case True
+    with k show ?thesis
+      apply (simp add: integrable_split)
+      apply (rule integral_unique [OF has_integral_split[OF _ _ k]])
+      apply (auto intro: integrable_integral)
+      done
+  next
+    case False
+    have "\<not> (f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<or> \<not> ( f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k})"
+    proof (rule ccontr)
+      assume "\<not> ?thesis"
+      then have "f integrable_on cbox a b"
+        unfolding integrable_on_def
+        apply (rule_tac x="integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f + integral (cbox a b \<inter> {x. x \<bullet> k \<ge> c}) f" in exI)
+        apply (rule has_integral_split[OF _ _ k])
+        apply (auto intro: integrable_integral)
+        done
+      then show False
+        using False by auto
+    qed
+    then show ?thesis
+      using False by auto
+  qed
+next
+  fix a b :: 'a
+  assume "content (cbox a b) = 0"
+  then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0"
+    using has_integral_null_eq
+    by (auto simp: integrable_on_null)
+qed
+
+subsection \<open>Finally, the integral of a constant\<close>
+
+lemma has_integral_const [intro]:
+  fixes a b :: "'a::euclidean_space"
+  shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)"
+  apply (auto intro!: exI [where x="\<lambda>x. ball x 1"] simp: split_def has_integral)
+  apply (subst scaleR_left.setsum[symmetric, unfolded o_def])
+  apply (subst additive_content_tagged_division[unfolded split_def])
+  apply auto
+  done
+
+lemma has_integral_const_real [intro]:
+  fixes a b :: real
+  shows "((\<lambda>x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}"
+  by (metis box_real(2) has_integral_const)
+
+lemma integral_const [simp]:
+  fixes a b :: "'a::euclidean_space"
+  shows "integral (cbox a b) (\<lambda>x. c) = content (cbox a b) *\<^sub>R c"
+  by (rule integral_unique) (rule has_integral_const)
+
+lemma integral_const_real [simp]:
+  fixes a b :: real
+  shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
+  by (metis box_real(2) integral_const)
+
+
+subsection \<open>Bounds on the norm of Riemann sums and the integral itself.\<close>
+
+lemma dsum_bound:
+  assumes "p division_of (cbox a b)"
+    and "norm c \<le> e"
+  shows "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)"
+proof -
+  have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = setsum content p"
+    apply (rule setsum.cong)
+    using assms
+    apply simp
+    apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4))
+    done
+  have e: "0 \<le> e"
+    using assms(2) norm_ge_zero order_trans by blast
+  have "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
+    using norm_setsum by blast
+  also have "...  \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)"
+    apply (simp add: setsum_right_distrib[symmetric] mult.commute)
+    using assms(2) mult_right_mono by blast
+  also have "... \<le> e * content (cbox a b)"
+    apply (rule mult_left_mono [OF _ e])
+    apply (simp add: sumeq)
+    using additive_content_division assms(1) eq_iff apply blast
+    done
+  finally show ?thesis .
+qed
+
+lemma rsum_bound:
+  assumes p: "p tagged_division_of (cbox a b)"
+      and "\<forall>x\<in>cbox a b. norm (f x) \<le> e"
+    shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content (cbox a b)"
+proof (cases "cbox a b = {}")
+  case True show ?thesis
+    using p unfolding True tagged_division_of_trivial by auto
+next
+  case False
+  then have e: "e \<ge> 0"
+    by (meson ex_in_conv assms(2) norm_ge_zero order_trans)
+  have setsum_le: "setsum (content \<circ> snd) p \<le> content (cbox a b)"
+    unfolding additive_content_tagged_division[OF p, symmetric] split_def
+    by (auto intro: eq_refl)
+  have con: "\<And>xk. xk \<in> p \<Longrightarrow> 0 \<le> content (snd xk)"
+    using tagged_division_ofD(4) [OF p] content_pos_le
+    by force
+  have norm: "\<And>xk. xk \<in> p \<Longrightarrow> norm (f (fst xk)) \<le> e"
+    unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms
+    by (metis prod.collapse subset_eq)
+  have "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))"
+    by (rule norm_setsum)
+  also have "...  \<le> e * content (cbox a b)"
+    unfolding split_def norm_scaleR
+    apply (rule order_trans[OF setsum_mono])
+    apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e])
+    apply (metis norm)
+    unfolding setsum_left_distrib[symmetric]
+    using con setsum_le
+    apply (auto simp: mult.commute intro: mult_left_mono [OF _ e])
+    done
+  finally show ?thesis .
+qed
+
+lemma rsum_diff_bound:
+  assumes "p tagged_division_of (cbox a b)"
+    and "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e"
+  shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le>
+         e * content (cbox a b)"
+  apply (rule order_trans[OF _ rsum_bound[OF assms]])
+  apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl)
+  done
+
+lemma has_integral_bound:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "0 \<le> B"
+      and "(f has_integral i) (cbox a b)"
+      and "\<forall>x\<in>cbox a b. norm (f x) \<le> B"
+    shows "norm i \<le> B * content (cbox a b)"
+proof (rule ccontr)
+  assume "\<not> ?thesis"
+  then have *: "norm i - B * content (cbox a b) > 0"
+    by auto
+  from assms(2)[unfolded has_integral,rule_format,OF *]
+  guess d by (elim exE conjE) note d=this[rule_format]
+  from fine_division_exists[OF this(1), of a b] guess p . note p=this
+  have *: "\<And>s B. norm s \<le> B \<Longrightarrow> \<not> norm (s - i) < norm i - B"
+    unfolding not_less
+    by (metis norm_triangle_sub[of i] add.commute le_less_trans less_diff_eq linorder_not_le norm_minus_commute)
+  show False
+    using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto
+qed
+
+corollary has_integral_bound_real:
+  fixes f :: "real \<Rightarrow> 'b::real_normed_vector"
+  assumes "0 \<le> B"
+      and "(f has_integral i) {a .. b}"
+      and "\<forall>x\<in>{a .. b}. norm (f x) \<le> B"
+    shows "norm i \<le> B * content {a .. b}"
+  by (metis assms box_real(2) has_integral_bound)
+
+corollary integrable_bound:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "0 \<le> B"
+      and "f integrable_on (cbox a b)"
+      and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B"
+    shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)"
+by (metis integrable_integral has_integral_bound assms)
+
+
+subsection \<open>Similar theorems about relationship among components.\<close>
+
+lemma rsum_component_le:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "p tagged_division_of (cbox a b)"
+      and "\<forall>x\<in>cbox a b. (f x)\<bullet>i \<le> (g x)\<bullet>i"
+    shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i"
+unfolding inner_setsum_left
+proof (rule setsum_mono, clarify)
+  fix a b
+  assume ab: "(a, b) \<in> p"
+  note tagged = tagged_division_ofD(2-4)[OF assms(1) ab]
+  from this(3) guess u v by (elim exE) note b=this
+  show "(content b *\<^sub>R f a) \<bullet> i \<le> (content b *\<^sub>R g a) \<bullet> i"
+    unfolding b inner_simps real_scaleR_def
+    apply (rule mult_left_mono)
+    using assms(2) tagged
+    by (auto simp add: content_pos_le)
+qed
+
+lemma has_integral_component_le:
+  fixes f g :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes k: "k \<in> Basis"
+  assumes "(f has_integral i) s" "(g has_integral j) s"
+    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
+  shows "i\<bullet>k \<le> j\<bullet>k"
+proof -
+  have lem: "i\<bullet>k \<le> j\<bullet>k"
+    if f_i: "(f has_integral i) (cbox a b)"
+    and g_j: "(g has_integral j) (cbox a b)"
+    and le: "\<forall>x\<in>cbox a b. (f x)\<bullet>k \<le> (g x)\<bullet>k"
+    for a b i and j :: 'b and f g :: "'a \<Rightarrow> 'b"
+  proof (rule ccontr)
+    assume "\<not> ?thesis"
+    then have *: "0 < (i\<bullet>k - j\<bullet>k) / 3"
+      by auto
+    guess d1 using f_i[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format]
+    guess d2 using g_j[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format]
+    obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p"
+       using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter
+       by metis
+    note le_less_trans[OF Basis_le_norm[OF k]]
+    then have "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
+              "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
+      using  k norm_bound_Basis_lt d1 d2 p
+      by blast+
+    then show False
+      unfolding inner_simps
+      using rsum_component_le[OF p(1) le]
+      by (simp add: abs_real_def split: if_split_asm)
+  qed
+  show ?thesis
+  proof (cases "\<exists>a b. s = cbox a b")
+    case True
+    with lem assms show ?thesis
+      by auto
+  next
+    case False
+    show ?thesis
+    proof (rule ccontr)
+      assume "\<not> i\<bullet>k \<le> j\<bullet>k"
+      then have ij: "(i\<bullet>k - j\<bullet>k) / 3 > 0"
+        by auto
+      note has_integral_altD[OF _ False this]
+      from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format]
+      have "bounded (ball 0 B1 \<union> ball (0::'a) B2)"
+        unfolding bounded_Un by(rule conjI bounded_ball)+
+      from bounded_subset_cbox[OF this] guess a b by (elim exE)
+      note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
+      guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this]
+      guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
+      have *: "\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False"
+        by (simp add: abs_real_def split: if_split_asm)
+      note le_less_trans[OF Basis_le_norm[OF k]]
+      note this[OF w1(2)] this[OF w2(2)]
+      moreover
+      have "w1\<bullet>k \<le> w2\<bullet>k"
+        by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4))
+      ultimately show False
+        unfolding inner_simps by(rule *)
+    qed
+  qed
+qed
+
+lemma integral_component_le:
+  fixes g f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "k \<in> Basis"
+    and "f integrable_on s" "g integrable_on s"
+    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
+  shows "(integral s f)\<bullet>k \<le> (integral s g)\<bullet>k"
+  apply (rule has_integral_component_le)
+  using integrable_integral assms
+  apply auto
+  done
+
+lemma has_integral_component_nonneg:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "k \<in> Basis"
+    and "(f has_integral i) s"
+    and "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k"
+  shows "0 \<le> i\<bullet>k"
+  using has_integral_component_le[OF assms(1) has_integral_0 assms(2)]
+  using assms(3-)
+  by auto
+
+lemma integral_component_nonneg:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "k \<in> Basis"
+    and  "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k"
+  shows "0 \<le> (integral s f)\<bullet>k"
+proof (cases "f integrable_on s")
+  case True show ?thesis
+    apply (rule has_integral_component_nonneg)
+    using assms True
+    apply auto
+    done
+next
+  case False then show ?thesis by (simp add: not_integrable_integral)
+qed
+
+lemma has_integral_component_neg:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "k \<in> Basis"
+    and "(f has_integral i) s"
+    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> 0"
+  shows "i\<bullet>k \<le> 0"
+  using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-)
+  by auto
+
+lemma has_integral_component_lbound:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "(f has_integral i) (cbox a b)"
+    and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k"
+    and "k \<in> Basis"
+  shows "B * content (cbox a b) \<le> i\<bullet>k"
+  using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\<Sum>i\<in>Basis. B *\<^sub>R i)::'b"] assms(2-)
+  by (auto simp add: field_simps)
+
+lemma has_integral_component_ubound:
+  fixes f::"'a::euclidean_space => 'b::euclidean_space"
+  assumes "(f has_integral i) (cbox a b)"
+    and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B"
+    and "k \<in> Basis"
+  shows "i\<bullet>k \<le> B * content (cbox a b)"
+  using has_integral_component_le[OF assms(3,1) has_integral_const, of "\<Sum>i\<in>Basis. B *\<^sub>R i"] assms(2-)
+  by (auto simp add: field_simps)
+
+lemma integral_component_lbound:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "f integrable_on cbox a b"
+    and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k"
+    and "k \<in> Basis"
+  shows "B * content (cbox a b) \<le> (integral(cbox a b) f)\<bullet>k"
+  apply (rule has_integral_component_lbound)
+  using assms
+  unfolding has_integral_integral
+  apply auto
+  done
+
+lemma integral_component_lbound_real:
+  assumes "f integrable_on {a ::real .. b}"
+    and "\<forall>x\<in>{a .. b}. B \<le> f(x)\<bullet>k"
+    and "k \<in> Basis"
+  shows "B * content {a .. b} \<le> (integral {a .. b} f)\<bullet>k"
+  using assms
+  by (metis box_real(2) integral_component_lbound)
+
+lemma integral_component_ubound:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes "f integrable_on cbox a b"
+    and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B"
+    and "k \<in> Basis"
+  shows "(integral (cbox a b) f)\<bullet>k \<le> B * content (cbox a b)"
+  apply (rule has_integral_component_ubound)
+  using assms
+  unfolding has_integral_integral
+  apply auto
+  done
+
+lemma integral_component_ubound_real:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "f integrable_on {a .. b}"
+    and "\<forall>x\<in>{a .. b}. f x\<bullet>k \<le> B"
+    and "k \<in> Basis"
+  shows "(integral {a .. b} f)\<bullet>k \<le> B * content {a .. b}"
+  using assms
+  by (metis box_real(2) integral_component_ubound)
+
+subsection \<open>Uniform limit of integrable functions is integrable.\<close>
+
+lemma real_arch_invD:
+  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
+  by (subst(asm) real_arch_inverse)
+
+lemma integrable_uniform_limit:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
+  shows "f integrable_on cbox a b"
+proof (cases "content (cbox a b) > 0")
+  case False then show ?thesis
+      using has_integral_null
+      by (simp add: content_lt_nz integrable_on_def)
+next
+  case True
+  have *: "\<And>P. \<forall>e>(0::real). P e \<Longrightarrow> \<forall>n::nat. P (inverse (real n + 1))"
+    by auto
+  from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format]
+  from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\<lambda>x. x"]]
+  obtain i where i: "\<And>x. (g x has_integral i x) (cbox a b)"
+      by auto
+  have "Cauchy i"
+    unfolding Cauchy_def
+  proof clarify
+    fix e :: real
+    assume "e>0"
+    then have "e / 4 / content (cbox a b) > 0"
+      using True by (auto simp add: field_simps)
+    then obtain M :: nat
+         where M: "M \<noteq> 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)"
+      by (subst (asm) real_arch_inverse) auto
+    show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (i m) (i n) < e"
+    proof (rule exI [where x=M], clarify)
+      fix m n
+      assume m: "M \<le> m" and n: "M \<le> n"
+      have "e/4>0" using \<open>e>0\<close> by auto
+      note * = i[unfolded has_integral,rule_format,OF this]
+      from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format]
+      from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format]
+      from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b]
+      obtain p where p: "p tagged_division_of cbox a b" "(\<lambda>x. gm x \<inter> gn x) fine p"
+        by auto
+      { fix s1 s2 i1 and i2::'b
+        assume no: "norm(s2 - s1) \<le> e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4"
+        have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
+          using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
+          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"]
+          by (auto simp add: algebra_simps)
+        also have "\<dots> < e"
+          using no
+          unfolding norm_minus_commute
+          by (auto simp add: algebra_simps)
+        finally have "norm (i1 - i2) < e" .
+      } note triangle3 = this
+      have finep: "gm fine p" "gn fine p"
+        using fine_inter p  by auto
+      { fix x
+        assume x: "x \<in> cbox a b"
+        have "norm (f x - g n x) + norm (f x - g m x) \<le> inverse (real n + 1) + inverse (real m + 1)"
+          using g(1)[OF x, of n] g(1)[OF x, of m] by auto
+        also have "\<dots> \<le> inverse (real M) + inverse (real M)"
+          apply (rule add_mono)
+          using M(2) m n by auto
+        also have "\<dots> = 2 / real M"
+          unfolding divide_inverse by auto
+        finally have "norm (g n x - g m x) \<le> 2 / real M"
+          using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
+          by (auto simp add: algebra_simps simp add: norm_minus_commute)
+      } note norm_le = this
+      have le_e2: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g n x) - (\<Sum>(x, k)\<in>p. content k *\<^sub>R g m x)) \<le> e / 2"
+        apply (rule order_trans [OF rsum_diff_bound[OF p(1), where e="2 / real M"]])
+        apply (blast intro: norm_le)
+        using M True
+        by (auto simp add: field_simps)
+      then show "dist (i m) (i n) < e"
+        unfolding dist_norm
+        using gm gn p finep
+        by (auto intro!: triangle3)
+    qed
+  qed
+  then obtain s where s: "i \<longlonglongrightarrow> s"
+    using convergent_eq_cauchy[symmetric] by blast
+  show ?thesis
+    unfolding integrable_on_def has_integral
+  proof (rule_tac x=s in exI, clarify)
+    fix e::real
+    assume e: "0 < e"
+    then have *: "e/3 > 0" by auto
+    then obtain N1 where N1: "\<forall>n\<ge>N1. norm (i n - s) < e / 3"
+      using LIMSEQ_D [OF s] by metis
+    from e True have "e / 3 / content (cbox a b) > 0"
+      by (auto simp add: field_simps)
+    from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this
+    from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format]
+    { fix sf sg i
+      assume no: "norm (sf - sg) \<le> e / 3"
+                 "norm(i - s) < e / 3"
+                 "norm (sg - i) < e / 3"
+      have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
+        using norm_triangle_ineq[of "sf - sg" "sg - s"]
+        using norm_triangle_ineq[of "sg -  i" " i - s"]
+        by (auto simp add: algebra_simps)
+      also have "\<dots> < e"
+        using no
+        unfolding norm_minus_commute
+        by (auto simp add: algebra_simps)
+      finally have "norm (sf - s) < e" .
+    } note lem = this
+    { fix p
+      assume p: "p tagged_division_of (cbox a b) \<and> g' fine p"
+      then have norm_less: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g (N1 + N2) x) - i (N1 + N2)) < e / 3"
+        using g' by blast
+      have "content (cbox a b) < e / 3 * (of_nat N2)"
+        using N2 unfolding inverse_eq_divide using True by (auto simp add: field_simps)
+      moreover have "e / 3 * of_nat N2 \<le> e / 3 * (of_nat (N1 + N2) + 1)"
+        using \<open>e>0\<close> by auto
+      ultimately have "content (cbox a b) < e / 3 * (of_nat (N1 + N2) + 1)"
+        by linarith
+      then have le_e3: "inverse (real (N1 + N2) + 1) * content (cbox a b) \<le> e / 3"
+        unfolding inverse_eq_divide
+        by (auto simp add: field_simps)
+      have ne3: "norm (i (N1 + N2) - s) < e / 3"
+        using N1 by auto
+      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e"
+        apply (rule lem[OF order_trans [OF _ le_e3] ne3 norm_less])
+        apply (rule rsum_diff_bound[OF p[THEN conjunct1]])
+        apply (blast intro: g)
+        done }
+    then show "\<exists>d. gauge d \<and>
+             (\<forall>p. p tagged_division_of cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e)"
+      by (blast intro: g')
+  qed
+qed
+
+lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified]
+
+
+subsection \<open>Negligible sets.\<close>
+
+definition "negligible (s:: 'a::euclidean_space set) \<longleftrightarrow>
+  (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) (cbox a b))"
+
+
+subsection \<open>Negligibility of hyperplane.\<close>
+
+lemma interval_doublesplit:
+  fixes a :: "'a::euclidean_space"
+  assumes "k \<in> Basis"
+  shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} =
+    cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i)
+     (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)"
+proof -
+  have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
+    by auto
+  have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}"
+    by blast
+  show ?thesis
+    unfolding * ** interval_split[OF assms] by (rule refl)
+qed
+
+lemma division_doublesplit:
+  fixes a :: "'a::euclidean_space"
+  assumes "p division_of (cbox a b)"
+    and k: "k \<in> Basis"
+  shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
+         division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
+proof -
+  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
+    by auto
+  have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
+    by auto
+  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
+  note division_split(2)[OF this, where c="c-e" and k=k,OF k]
+  then show ?thesis
+    apply (rule **)
+    subgoal
+      apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric])
+      apply (rule equalityI)
+      apply blast
+      apply clarsimp
+      apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
+      apply auto
+      done
+    by (simp add: interval_split k interval_doublesplit)
+qed
+
+lemma content_doublesplit:
+  fixes a :: "'a::euclidean_space"
+  assumes "0 < e"
+    and k: "k \<in> Basis"
+  obtains d where "0 < d" and "content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) < e"
+proof (cases "content (cbox a b) = 0")
+  case True
+  then have ce: "content (cbox a b) < e"
+    by (metis \<open>0 < e\<close>)
+  show ?thesis
+    apply (rule that[of 1])
+    apply simp
+    unfolding interval_doublesplit[OF k]
+    apply (rule le_less_trans[OF content_subset ce])
+    apply (auto simp: interval_doublesplit[symmetric] k)
+    done
+next
+  case False
+  define d where "d = e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
+  note False[unfolded content_eq_0 not_ex not_le, rule_format]
+  then have "\<And>x. x \<in> Basis \<Longrightarrow> b\<bullet>x > a\<bullet>x"
+    by (auto simp add:not_le)
+  then have prod0: "0 < setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
+    by (force simp add: setprod_pos field_simps)
+  then have "d > 0"
+    using assms
+    by (auto simp add: d_def field_simps)
+  then show ?thesis
+  proof (rule that[of d])
+    have *: "Basis = insert k (Basis - {k})"
+      using k by auto
+    have less_e: "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e"
+    proof -
+      have "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) \<le> 2 * d"
+        by auto
+      also have "\<dots> < e / (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i)"
+        unfolding d_def
+        using assms prod0
+        by (auto simp add: field_simps)
+      finally show "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e"
+        unfolding pos_less_divide_eq[OF prod0] .
+    qed
+    show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
+    proof (cases "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = {}")
+      case True
+      then show ?thesis
+        using assms by simp
+    next
+      case False
+      then have
+          "(\<Prod>i\<in>Basis - {k}. interval_upperbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i -
+                interval_lowerbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i)
+           = (\<Prod>i\<in>Basis - {k}. b\<bullet>i - a\<bullet>i)"
+        by (simp add: box_eq_empty interval_doublesplit[OF k])
+      then show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
+        unfolding content_def
+        using assms False
+        apply (subst *)
+        apply (subst setprod.insert)
+        apply (simp_all add: interval_doublesplit[OF k] box_eq_empty not_less less_e)
+        done
+    qed
+  qed
+qed
+
+lemma negligible_standard_hyperplane[intro]:
+  fixes k :: "'a::euclidean_space"
+  assumes k: "k \<in> Basis"
+  shows "negligible {x. x\<bullet>k = c}"
+  unfolding negligible_def has_integral
+proof (clarify, goal_cases)
+  case (1 a b e)
+  from this and k obtain d where d: "0 < d" "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
+    by (rule content_doublesplit)
+  let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real"
+  show ?case
+    apply (rule_tac x="\<lambda>x. ball x d" in exI)
+    apply rule
+    apply (rule gauge_ball)
+    apply (rule d)
+  proof (rule, rule)
+    fix p
+    assume p: "p tagged_division_of (cbox a b) \<and> (\<lambda>x. ball x d) fine p"
+    have *: "(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) =
+      (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)"
+      apply (rule setsum.cong)
+      apply (rule refl)
+      unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
+      apply cases
+      apply (rule disjI1)
+      apply assumption
+      apply (rule disjI2)
+    proof -
+      fix x l
+      assume as: "(x, l) \<in> p" "?i x \<noteq> 0"
+      then have xk: "x\<bullet>k = c"
+        unfolding indicator_def
+        apply -
+        apply (rule ccontr)
+        apply auto
+        done
+      show "content l = content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
+        apply (rule arg_cong[where f=content])
+        apply (rule set_eqI)
+        apply rule
+        apply rule
+        unfolding mem_Collect_eq
+      proof -
+        fix y
+        assume y: "y \<in> l"
+        note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
+        note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y]
+        note le_less_trans[OF Basis_le_norm[OF k] this]
+        then show "\<bar>y \<bullet> k - c\<bar> \<le> d"
+          unfolding inner_simps xk by auto
+      qed auto
+    qed
+    note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
+    show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e"
+      unfolding diff_0_right *
+      unfolding real_scaleR_def real_norm_def
+      apply (subst abs_of_nonneg)
+      apply (rule setsum_nonneg)
+      apply rule
+      unfolding split_paired_all split_conv
+      apply (rule mult_nonneg_nonneg)
+      apply (drule p'(4))
+      apply (erule exE)+
+      apply(rule_tac b=b in back_subst)
+      prefer 2
+      apply (subst(asm) eq_commute)
+      apply assumption
+      apply (subst interval_doublesplit[OF k])
+      apply (rule content_pos_le)
+      apply (rule indicator_pos_le)
+    proof -
+      have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le>
+        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
+        apply (rule setsum_mono)
+        unfolding split_paired_all split_conv
+        apply (rule mult_right_le_one_le)
+        apply (drule p'(4))
+        apply (auto simp add:interval_doublesplit[OF k])
+        done
+      also have "\<dots> < e"
+      proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases)
+        case prems: (1 u v)
+        have "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content (cbox u v)"
+          unfolding interval_doublesplit[OF k]
+          apply (rule content_subset)
+          unfolding interval_doublesplit[symmetric,OF k]
+          apply auto
+          done
+        then show ?case
+          unfolding prems interval_doublesplit[OF k]
+          by (blast intro: antisym)
+      next
+        have "(\<Sum>l\<in>snd ` p. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
+          setsum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
+        proof (subst (2) setsum.reindex_nontrivial)
+          fix x y assume "x \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}"
+            "x \<noteq> y" and eq: "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
+          then obtain x' y' where "(x', x) \<in> p" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> p" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}"
+            by (auto)
+          from p'(5)[OF \<open>(x', x) \<in> p\<close> \<open>(y', y) \<in> p\<close>] \<open>x \<noteq> y\<close> have "interior (x \<inter> y) = {}"
+            by auto
+          moreover have "interior ((x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<inter> (y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<subseteq> interior (x \<inter> y)"
+            by (auto intro: interior_mono)
+          ultimately have "interior (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = {}"
+            by (auto simp: eq)
+          then show "content (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
+            using p'(4)[OF \<open>(x', x) \<in> p\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int)
+        qed (insert p'(1), auto intro!: setsum.mono_neutral_right)
+        also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}. content l *\<^sub>R 1::real)"
+          by simp
+        also have "\<dots> \<le> 1 * content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
+          using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]]
+          unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto
+        also have "\<dots> < e"
+          using d(2) by simp
+        finally show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e" .
+      qed
+      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" .
+    qed
+  qed
+qed
+
+
+subsection \<open>A technical lemma about "refinement" of division.\<close>
+
+lemma tagged_division_finer:
+  fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
+  assumes "p tagged_division_of (cbox a b)"
+    and "gauge d"
+  obtains q where "q tagged_division_of (cbox a b)"
+    and "d fine q"
+    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
+proof -
+  let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow>
+    (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and>
+      (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))"
+  {
+    have *: "finite p" "p tagged_partial_division_of (cbox a b)"
+      using assms(1)
+      unfolding tagged_division_of_def
+      by auto
+    presume "\<And>p. finite p \<Longrightarrow> ?P p"
+    from this[rule_format,OF * assms(2)] guess q .. note q=this
+    then show ?thesis
+      apply -
+      apply (rule that[of q])
+      unfolding tagged_division_ofD[OF assms(1)]
+      apply auto
+      done
+  }
+  fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
+  assume as: "finite p"
+  show "?P p"
+    apply rule
+    apply rule
+    using as
+  proof (induct p)
+    case empty
+    show ?case
+      apply (rule_tac x="{}" in exI)
+      unfolding fine_def
+      apply auto
+      done
+  next
+    case (insert xk p)
+    guess x k using surj_pair[of xk] by (elim exE) note xk=this
+    note tagged_partial_division_subset[OF insert(4) subset_insertI]
+    from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this]
+    have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
+      unfolding xk by auto
+    note p = tagged_partial_division_ofD[OF insert(4)]
+    from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this
+
+    have "finite {k. \<exists>x. (x, k) \<in> p}"
+      apply (rule finite_subset[of _ "snd ` p"])
+      using p
+      apply safe
+      apply (metis image_iff snd_conv)
+      apply auto
+      done
+    then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
+      apply (rule inter_interior_unions_intervals)
+      apply (rule open_interior)
+      apply (rule_tac[!] ballI)
+      unfolding mem_Collect_eq
+      apply (erule_tac[!] exE)
+      apply (drule p(4)[OF insertI2])
+      apply assumption
+      apply (rule p(5))
+      unfolding uv xk
+      apply (rule insertI1)
+      apply (rule insertI2)
+      apply assumption
+      using insert(2)
+      unfolding uv xk
+      apply auto
+      done
+    show ?case
+    proof (cases "cbox u v \<subseteq> d x")
+      case True
+      then show ?thesis
+        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
+        apply rule
+        unfolding * uv
+        apply (rule tagged_division_union)
+        apply (rule tagged_division_of_self)
+        apply (rule p[unfolded xk uv] insertI1)+
+        apply (rule q1)
+        apply (rule int)
+        apply rule
+        apply (rule fine_union)
+        apply (subst fine_def)
+        defer
+        apply (rule q1)
+        unfolding Ball_def split_paired_All split_conv
+        apply rule
+        apply rule
+        apply rule
+        apply rule
+        apply (erule insertE)
+        apply (simp add: uv xk)
+        apply (rule UnI2)
+        apply (drule q1(3)[rule_format])
+        unfolding xk uv
+        apply auto
+        done
+    next
+      case False
+      from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this
+      show ?thesis
+        apply (rule_tac x="q2 \<union> q1" in exI)
+        apply rule
+        unfolding * uv
+        apply (rule tagged_division_union q2 q1 int fine_union)+
+        unfolding Ball_def split_paired_All split_conv
+        apply rule
+        apply (rule fine_union)
+        apply (rule q1 q2)+
+        apply rule
+        apply rule
+        apply rule
+        apply rule
+        apply (erule insertE)
+        apply (rule UnI2)
+        apply (simp add: False uv xk)
+        apply (drule q1(3)[rule_format])
+        using False
+        unfolding xk uv
+        apply auto
+        done
+    qed
+  qed
+qed
+
+
+subsection \<open>Hence the main theorem about negligible sets.\<close>
+
+lemma finite_product_dependent:
+  assumes "finite s"
+    and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)"
+  shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
+  using assms
+proof induct
+  case (insert x s)
+  have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} =
+    (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
+  show ?case
+    unfolding *
+    apply (rule finite_UnI)
+    using insert
+    apply auto
+    done
+qed auto
+
+lemma sum_sum_product:
+  assumes "finite s"
+    and "\<forall>i\<in>s. finite (t i)"
+  shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s =
+    setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
+  using assms
+proof induct
+  case (insert a s)
+  have *: "{(i, j) |i j. i \<in> insert a s \<and> j \<in> t i} =
+    (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
+  show ?case
+    unfolding *
+    apply (subst setsum.union_disjoint)
+    unfolding setsum.insert[OF insert(1-2)]
+    prefer 4
+    apply (subst insert(3))
+    unfolding add_right_cancel
+  proof -
+    show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
+      apply (subst setsum.reindex)
+      unfolding inj_on_def
+      apply auto
+      done
+    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
+      apply (rule finite_product_dependent)
+      using insert
+      apply auto
+      done
+  qed (insert insert, auto)
+qed auto
+
+lemma has_integral_negligible:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "negligible s"
+    and "\<forall>x\<in>(t - s). f x = 0"
+  shows "(f has_integral 0) t"
+proof -
+  presume P: "\<And>f::'b::euclidean_space \<Rightarrow> 'a.
+    \<And>a b. \<forall>x. x \<notin> s \<longrightarrow> f x = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
+  let ?f = "(\<lambda>x. if x \<in> t then f x else 0)"
+  show ?thesis
+    apply (rule_tac f="?f" in has_integral_eq)
+    unfolding if_P
+    apply (rule refl)
+    apply (subst has_integral_alt)
+    apply cases
+    apply (subst if_P, assumption)
+    unfolding if_not_P
+  proof -
+    assume "\<exists>a b. t = cbox a b"
+    then guess a b apply - by (erule exE)+ note t = this
+    show "(?f has_integral 0) t"
+      unfolding t
+      apply (rule P)
+      using assms(2)
+      unfolding t
+      apply auto
+      done
+  next
+    show "\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> t then ?f x else 0) has_integral z) (cbox a b) \<and> norm (z - 0) < e)"
+      apply safe
+      apply (rule_tac x=1 in exI)
+      apply rule
+      apply (rule zero_less_one)
+      apply safe
+      apply (rule_tac x=0 in exI)
+      apply rule
+      apply (rule P)
+      using assms(2)
+      apply auto
+      done
+  qed
+next
+  fix f :: "'b \<Rightarrow> 'a"
+  fix a b :: 'b
+  assume assm: "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
+  show "(f has_integral 0) (cbox a b)"
+    unfolding has_integral
+  proof (safe, goal_cases)
+    case prems: (1 e)
+    then have "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0"
+      apply -
+      apply (rule divide_pos_pos)
+      defer
+      apply (rule mult_pos_pos)
+      apply (auto simp add:field_simps)
+      done
+    note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b]
+    note allI[OF this,of "\<lambda>x. x"]
+    from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]]
+    show ?case
+      apply (rule_tac x="\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x" in exI)
+    proof safe
+      show "gauge (\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x)"
+        using d(1) unfolding gauge_def by auto
+      fix p
+      assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x) fine p"
+      let ?goal = "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
+      {
+        presume "p \<noteq> {} \<Longrightarrow> ?goal"
+        then show ?goal
+          apply (cases "p = {}")
+          using prems
+          apply auto
+          done
+      }
+      assume as': "p \<noteq> {}"
+      from real_arch_simple[of "Max((\<lambda>(x,k). norm(f x)) ` p)"] guess N ..
+      then have N: "\<forall>x\<in>(\<lambda>(x, k). norm (f x)) ` p. x \<le> real N"
+        by (meson Max_ge as(1) dual_order.trans finite_imageI tagged_division_of_finite)
+      have "\<forall>i. \<exists>q. q tagged_division_of (cbox a b) \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)"
+        by (auto intro: tagged_division_finer[OF as(1) d(1)])
+      from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
+      have *: "\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> (0::real)"
+        apply (rule setsum_nonneg)
+        apply safe
+        unfolding real_scaleR_def
+        apply (drule tagged_division_ofD(4)[OF q(1)])
+        apply (auto intro: mult_nonneg_nonneg)
+        done
+      have **: "finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow>
+        (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t" for f g s t
+        apply (rule setsum_le_included[of s t g snd f])
+        prefer 4
+        apply safe
+        apply (erule_tac x=x in ballE)
+        apply (erule exE)
+        apply (rule_tac x="(xa,x)" in bexI)
+        apply auto
+        done
+      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
+        norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}"
+        unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
+        apply (rule order_trans)
+        apply (rule norm_setsum)
+        apply (subst sum_sum_product)
+        prefer 3
+      proof (rule **, safe)
+        show "finite {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i}"
+          apply (rule finite_product_dependent)
+          using q
+          apply auto
+          done
+        fix i a b
+        assume as'': "(a, b) \<in> q i"
+        show "0 \<le> (real i + 1) * (content b *\<^sub>R indicator s a)"
+          unfolding real_scaleR_def
+          using tagged_division_ofD(4)[OF q(1) as'']
+          by (auto intro!: mult_nonneg_nonneg)
+      next
+        fix i :: nat
+        show "finite (q i)"
+          using q by auto
+      next
+        fix x k
+        assume xk: "(x, k) \<in> p"
+        define n where "n = nat \<lfloor>norm (f x)\<rfloor>"
+        have *: "norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p"
+          using xk by auto
+        have nfx: "real n \<le> norm (f x)" "norm (f x) \<le> real n + 1"
+          unfolding n_def by auto
+        then have "n \<in> {0..N + 1}"
+          using N[rule_format,OF *] by auto
+        moreover
+        note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv]
+        note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this]
+        note this[unfolded n_def[symmetric]]
+        moreover
+        have "norm (content k *\<^sub>R f x) \<le> (real n + 1) * (content k * indicator s x)"
+        proof (cases "x \<in> s")
+          case False
+          then show ?thesis
+            using assm by auto
+        next
+          case True
+          have *: "content k \<ge> 0"
+            using tagged_division_ofD(4)[OF as(1) xk] by auto
+          moreover
+          have "content k * norm (f x) \<le> content k * (real n + 1)"
+            apply (rule mult_mono)
+            using nfx *
+            apply auto
+            done
+          ultimately
+          show ?thesis
+            unfolding abs_mult
+            using nfx True
+            by (auto simp add: field_simps)
+        qed
+        ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le>
+          (real y + 1) * (content k *\<^sub>R indicator s x)"
+          apply (rule_tac x=n in exI)
+          apply safe
+          apply (rule_tac x=n in exI)
+          apply (rule_tac x="(x,k)" in exI)
+          apply safe
+          apply auto
+          done
+      qed (insert as, auto)
+      also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}"
+      proof (rule setsum_mono, goal_cases)
+        case (1 i)
+        then show ?case
+          apply (subst mult.commute, subst pos_le_divide_eq[symmetric])
+          using d(2)[rule_format, of "q i" i]
+          using q[rule_format]
+          apply (auto simp add: field_simps)
+          done
+      qed
+      also have "\<dots> < e * inverse 2 * 2"
+        unfolding divide_inverse setsum_right_distrib[symmetric]
+        apply (rule mult_strict_left_mono)
+        unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric]
+        apply (subst geometric_sum)
+        using prems
+        apply auto
+        done
+      finally show "?goal" by auto
+    qed
+  qed
+qed
+
+lemma has_integral_spike:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
+  assumes "negligible s"
+    and "(\<forall>x\<in>(t - s). g x = f x)"
+    and "(f has_integral y) t"
+  shows "(g has_integral y) t"
+proof -
+  {
+    fix a b :: 'b
+    fix f g :: "'b \<Rightarrow> 'a"
+    fix y :: 'a
+    assume as: "\<forall>x \<in> cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)"
+    have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)"
+      apply (rule has_integral_add[OF as(2)])
+      apply (rule has_integral_negligible[OF assms(1)])
+      using as
+      apply auto
+      done
+    then have "(g has_integral y) (cbox a b)"
+      by auto
+  } note * = this
+  show ?thesis
+    apply (subst has_integral_alt)
+    using assms(2-)
+    apply -
+    apply (rule cond_cases)
+    apply safe
+    apply (rule *)
+    apply assumption+
+    apply (subst(asm) has_integral_alt)
+    unfolding if_not_P
+    apply (erule_tac x=e in allE)
+    apply safe
+    apply (rule_tac x=B in exI)
+    apply safe
+    apply (erule_tac x=a in allE)
+    apply (erule_tac x=b in allE)
+    apply safe
+    apply (rule_tac x=z in exI)
+    apply safe
+    apply (rule *[where fa2="\<lambda>x. if x\<in>t then f x else 0"])
+    apply auto
+    done
+qed
+
+lemma has_integral_spike_eq:
+  assumes "negligible s"
+    and "\<forall>x\<in>(t - s). g x = f x"
+  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
+  apply rule
+  apply (rule_tac[!] has_integral_spike[OF assms(1)])
+  using assms(2)
+  apply auto
+  done
+
+lemma integrable_spike:
+  assumes "negligible s"
+    and "\<forall>x\<in>(t - s). g x = f x"
+    and "f integrable_on t"
+  shows "g integrable_on  t"
+  using assms
+  unfolding integrable_on_def
+  apply -
+  apply (erule exE)
+  apply rule
+  apply (rule has_integral_spike)
+  apply fastforce+
+  done
+
+lemma integral_spike:
+  assumes "negligible s"
+    and "\<forall>x\<in>(t - s). g x = f x"
+  shows "integral t f = integral t g"
+  using has_integral_spike_eq[OF assms] by (simp add: integral_def integrable_on_def)
+
+
+subsection \<open>Some other trivialities about negligible sets.\<close>
+
+lemma negligible_subset[intro]:
+  assumes "negligible s"
+    and "t \<subseteq> s"
+  shows "negligible t"
+  unfolding negligible_def
+proof (safe, goal_cases)
+  case (1 a b)
+  show ?case
+    using assms(1)[unfolded negligible_def,rule_format,of a b]
+    apply -
+    apply (rule has_integral_spike[OF assms(1)])
+    defer
+    apply assumption
+    using assms(2)
+    unfolding indicator_def
+    apply auto
+    done
+qed
+
+lemma negligible_diff[intro?]:
+  assumes "negligible s"
+  shows "negligible (s - t)"
+  using assms by auto
+
+lemma negligible_Int:
+  assumes "negligible s \<or> negligible t"
+  shows "negligible (s \<inter> t)"
+  using assms by auto
+
+lemma negligible_Un:
+  assumes "negligible s"
+    and "negligible t"
+  shows "negligible (s \<union> t)"
+  unfolding negligible_def
+proof (safe, goal_cases)
+  case (1 a b)
+  note assm = assms[unfolded negligible_def,rule_format,of a b]
+  then show ?case
+    apply (subst has_integral_spike_eq[OF assms(2)])
+    defer
+    apply assumption
+    unfolding indicator_def
+    apply auto
+    done
+qed
+
+lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t"
+  using negligible_Un by auto
+
+lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}"
+  using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] by auto
+
+lemma negligible_insert[simp]: "negligible (insert a s) \<longleftrightarrow> negligible s"
+  apply (subst insert_is_Un)
+  unfolding negligible_Un_eq
+  apply auto
+  done
+
+lemma negligible_empty[iff]: "negligible {}"
+  by auto
+
+lemma negligible_finite[intro]:
+  assumes "finite s"
+  shows "negligible s"
+  using assms by (induct s) auto
+
+lemma negligible_Union[intro]:
+  assumes "finite s"
+    and "\<forall>t\<in>s. negligible t"
+  shows "negligible(\<Union>s)"
+  using assms by induct auto
+
+lemma negligible:
+  "negligible s \<longleftrightarrow> (\<forall>t::('a::euclidean_space) set. ((indicator s::'a\<Rightarrow>real) has_integral 0) t)"
+  apply safe
+  defer
+  apply (subst negligible_def)
+proof -
+  fix t :: "'a set"
+  assume as: "negligible s"
+  have *: "(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)"
+    by auto
+  show "((indicator s::'a\<Rightarrow>real) has_integral 0) t"
+    apply (subst has_integral_alt)
+    apply cases
+    apply (subst if_P,assumption)
+    unfolding if_not_P
+    apply safe
+    apply (rule as[unfolded negligible_def,rule_format])
+    apply (rule_tac x=1 in exI)
+    apply safe
+    apply (rule zero_less_one)
+    apply (rule_tac x=0 in exI)
+    using negligible_subset[OF as,of "s \<inter> t"]
+    unfolding negligible_def indicator_def [abs_def]
+    unfolding *
+    apply auto
+    done
+qed auto
+
+
+subsection \<open>Finite case of the spike theorem is quite commonly needed.\<close>
+
+lemma has_integral_spike_finite:
+  assumes "finite s"
+    and "\<forall>x\<in>t-s. g x = f x"
+    and "(f has_integral y) t"
+  shows "(g has_integral y) t"
+  apply (rule has_integral_spike)
+  using assms
+  apply auto
+  done
+
+lemma has_integral_spike_finite_eq:
+  assumes "finite s"
+    and "\<forall>x\<in>t-s. g x = f x"
+  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
+  apply rule
+  apply (rule_tac[!] has_integral_spike_finite)
+  using assms
+  apply auto
+  done
+
+lemma integrable_spike_finite:
+  assumes "finite s"
+    and "\<forall>x\<in>t-s. g x = f x"
+    and "f integrable_on t"
+  shows "g integrable_on  t"
+  using assms
+  unfolding integrable_on_def
+  apply safe
+  apply (rule_tac x=y in exI)
+  apply (rule has_integral_spike_finite)
+  apply auto
+  done
+
+
+subsection \<open>In particular, the boundary of an interval is negligible.\<close>
+
+lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)"
+proof -
+  let ?A = "\<Union>((\<lambda>k. {x. x\<bullet>k = a\<bullet>k} \<union> {x::'a. x\<bullet>k = b\<bullet>k}) ` Basis)"
+  have "cbox a b - box a b \<subseteq> ?A"
+    apply rule unfolding Diff_iff mem_box
+    apply simp
+    apply(erule conjE bexE)+
+    apply(rule_tac x=i in bexI)
+    apply auto
+    done
+  then show ?thesis
+    apply -
+    apply (rule negligible_subset[of ?A])
+    apply (rule negligible_Union[OF finite_imageI])
+    apply auto
+    done
+qed
+
+lemma has_integral_spike_interior:
+  assumes "\<forall>x\<in>box a b. g x = f x"
+    and "(f has_integral y) (cbox a b)"
+  shows "(g has_integral y) (cbox a b)"
+  apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)])
+  using assms(1)
+  apply auto
+  done
+
+lemma has_integral_spike_interior_eq:
+  assumes "\<forall>x\<in>box a b. g x = f x"
+  shows "(f has_integral y) (cbox a b) \<longleftrightarrow> (g has_integral y) (cbox a b)"
+  apply rule
+  apply (rule_tac[!] has_integral_spike_interior)
+  using assms
+  apply auto
+  done
+
+lemma integrable_spike_interior:
+  assumes "\<forall>x\<in>box a b. g x = f x"
+    and "f integrable_on cbox a b"
+  shows "g integrable_on cbox a b"
+  using assms
+  unfolding integrable_on_def
+  using has_integral_spike_interior[OF assms(1)]
+  by auto
+
+
+subsection \<open>Integrability of continuous functions.\<close>
+
+lemma operative_approximable:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "0 \<le> e"
+  shows "comm_monoid.operative op \<and> True (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)"
+  unfolding comm_monoid.operative_def[OF comm_monoid_and]
+proof safe
+  fix a b :: 'b
+  show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
+    if "content (cbox a b) = 0"
+    apply (rule_tac x=f in exI)
+    using assms that
+    apply (auto intro!: integrable_on_null)
+    done
+  {
+    fix c g
+    fix k :: 'b
+    assume as: "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b"
+    assume k: "k \<in> Basis"
+    show "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
+      "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
+      apply (rule_tac[!] x=g in exI)
+      using as(1) integrable_split[OF as(2) k]
+      apply auto
+      done
+  }
+  fix c k g1 g2
+  assume as: "\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
+    "\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e" "g2 integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
+  assume k: "k \<in> Basis"
+  let ?g = "\<lambda>x. if x\<bullet>k = c then f x else if x\<bullet>k \<le> c then g1 x else g2 x"
+  show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
+    apply (rule_tac x="?g" in exI)
+    apply safe
+  proof goal_cases
+    case (1 x)
+    then show ?case
+      apply -
+      apply (cases "x\<bullet>k=c")
+      apply (case_tac "x\<bullet>k < c")
+      using as assms
+      apply auto
+      done
+  next
+    case 2
+    presume "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
+      and "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
+    then guess h1 h2 unfolding integrable_on_def by auto
+    from has_integral_split[OF this k] show ?case
+      unfolding integrable_on_def by auto
+  next
+    show "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
+      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]])
+      using k as(2,4)
+      apply auto
+      done
+  qed
+qed
+
+lemma comm_monoid_set_F_and: "comm_monoid_set.F op \<and> True f s \<longleftrightarrow> (finite s \<longrightarrow> (\<forall>x\<in>s. f x))"
+proof -
+  interpret bool: comm_monoid_set "op \<and>" True
+    proof qed auto
+  show ?thesis
+    by (induction s rule: infinite_finite_induct) auto
+qed
+
+lemma approximable_on_division:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "0 \<le> e"
+    and "d division_of (cbox a b)"
+    and "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
+  obtains g where "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b"
+proof -
+  note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_approximable[OF assms(1)] assms(2)]
+  from assms(3) this[unfolded comm_monoid_set_F_and, of f] division_of_finite[OF assms(2)]
+  guess g by auto
+  then show thesis
+    apply -
+    apply (rule that[of g])
+    apply auto
+    done
+qed
+
+lemma integrable_continuous:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "continuous_on (cbox a b) f"
+  shows "f integrable_on cbox a b"
+proof (rule integrable_uniform_limit, safe)
+  fix e :: real
+  assume e: "e > 0"
+  from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d ..
+  note d=conjunctD2[OF this,rule_format]
+  from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this
+  note p' = tagged_division_ofD[OF p(1)]
+  have *: "\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
+  proof (safe, unfold snd_conv)
+    fix x l
+    assume as: "(x, l) \<in> p"
+    from p'(4)[OF this] guess a b by (elim exE) note l=this
+    show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l"
+      apply (rule_tac x="\<lambda>y. f x" in exI)
+    proof safe
+      show "(\<lambda>y. f x) integrable_on l"
+        unfolding integrable_on_def l
+        apply rule
+        apply (rule has_integral_const)
+        done
+      fix y
+      assume y: "y \<in> l"
+      note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this]
+      note d(2)[OF _ _ this[unfolded mem_ball]]
+      then show "norm (f y - f x) \<le> e"
+        using y p'(2-3)[OF as] unfolding dist_norm l norm_minus_commute by fastforce
+    qed
+  qed
+  from e have "e \<ge> 0"
+    by auto
+  from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g .
+  then show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
+    by auto
+qed
+
+lemma integrable_continuous_real:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a .. b} f"
+  shows "f integrable_on {a .. b}"
+  by (metis assms box_real(2) integrable_continuous)
+
+subsection \<open>Specialization of additivity to one dimension.\<close>
+
+subsection \<open>Special case of additivity we need for the FTC.\<close>
+
+lemma additive_tagged_division_1:
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b"
+    and "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
+proof -
+  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
+  have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
+    using assms by auto
+  have *: "add.operative ?f"
+    unfolding add.operative_1_lt box_eq_empty
+    by auto
+  have **: "cbox a b \<noteq> {}"
+    using assms(1) by auto
+  note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
+  note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric]
+  show ?thesis
+    unfolding *
+    apply (rule setsum.cong)
+    unfolding split_paired_all split_conv
+    using assms(2)
+    apply auto
+    done
+qed
+
+
+subsection \<open>A useful lemma allowing us to factor out the content size.\<close>
+
+lemma has_integral_factor_content:
+  "(f has_integral i) (cbox a b) \<longleftrightarrow>
+    (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))"
+proof (cases "content (cbox a b) = 0")
+  case True
+  show ?thesis
+    unfolding has_integral_null_eq[OF True]
+    apply safe
+    apply (rule, rule, rule gauge_trivial, safe)
+    unfolding setsum_content_null[OF True] True
+    defer
+    apply (erule_tac x=1 in allE)
+    apply safe
+    defer
+    apply (rule fine_division_exists[of _ a b])
+    apply assumption
+    apply (erule_tac x=p in allE)
+    unfolding setsum_content_null[OF True]
+    apply auto
+    done
+next
+  case False
+  note F = this[unfolded content_lt_nz[symmetric]]
+  let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and>
+    (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)"
+  show ?thesis
+    apply (subst has_integral)
+  proof safe
+    fix e :: real
+    assume e: "e > 0"
+    {
+      assume "\<forall>e>0. ?P e op <"
+      then show "?P (e * content (cbox a b)) op \<le>"
+        apply (erule_tac x="e * content (cbox a b)" in allE)
+        apply (erule impE)
+        defer
+        apply (erule exE,rule_tac x=d in exI)
+        using F e
+        apply (auto simp add:field_simps)
+        done
+    }
+    {
+      assume "\<forall>e>0. ?P (e * content (cbox a b)) op \<le>"
+      then show "?P e op <"
+        apply (erule_tac x="e / 2 / content (cbox a b)" in allE)
+        apply (erule impE)
+        defer
+        apply (erule exE,rule_tac x=d in exI)
+        using F e
+        apply (auto simp add: field_simps)
+        done
+    }
+  qed
+qed
+
+lemma has_integral_factor_content_real:
+  "(f has_integral i) {a .. b::real} \<longleftrightarrow>
+    (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b}  \<and> d fine p \<longrightarrow>
+      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a .. b} ))"
+  unfolding box_real[symmetric]
+  by (rule has_integral_factor_content)
+
+
+subsection \<open>Fundamental theorem of calculus.\<close>
+
+lemma interval_bounds_real:
+  fixes q b :: real
+  assumes "a \<le> b"
+  shows "Sup {a..b} = b"
+    and "Inf {a..b} = a"
+  using assms by auto
+
+lemma fundamental_theorem_of_calculus:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "a \<le> b"
+    and "\<forall>x\<in>{a .. b}. (f has_vector_derivative f' x) (at x within {a .. b})"
+  shows "(f' has_integral (f b - f a)) {a .. b}"
+  unfolding has_integral_factor_content box_real[symmetric]
+proof safe
+  fix e :: real
+  assume e: "e > 0"
+  note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt]
+  have *: "\<And>P Q. \<forall>x\<in>{a .. b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a .. b} \<longrightarrow> Q x e d"
+    using e by blast
+  note this[OF assm,unfolded gauge_existence_lemma]
+  from choice[OF this,unfolded Ball_def[symmetric]] guess d ..
+  note d=conjunctD2[OF this[rule_format],rule_format]
+  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b))"
+    apply (rule_tac x="\<lambda>x. ball x (d x)" in exI)
+    apply safe
+    apply (rule gauge_ball_dependent)
+    apply rule
+    apply (rule d(1))
+  proof -
+    fix p
+    assume as: "p tagged_division_of cbox a b" "(\<lambda>x. ball x (d x)) fine p"
+    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)"
+      unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric]
+      unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\<lambda>x. x",symmetric]
+      unfolding setsum_right_distrib
+      defer
+      unfolding setsum_subtractf[symmetric]
+    proof (rule setsum_norm_le,safe)
+      fix x k
+      assume "(x, k) \<in> p"
+      note xk = tagged_division_ofD(2-4)[OF as(1) this]
+      from this(3) guess u v by (elim exE) note k=this
+      have *: "u \<le> v"
+        using xk unfolding k by auto
+      have ball: "\<forall>xa\<in>k. xa \<in> ball x (d x)"
+        using as(2)[unfolded fine_def,rule_format,OF \<open>(x,k)\<in>p\<close>,unfolded split_conv subset_eq] .
+      have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le>
+        norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)"
+        apply (rule order_trans[OF _ norm_triangle_ineq4])
+        apply (rule eq_refl)
+        apply (rule arg_cong[where f=norm])
+        unfolding scaleR_diff_left
+        apply (auto simp add:algebra_simps)
+        done
+      also have "\<dots> \<le> e * norm (u - x) + e * norm (v - x)"
+        apply (rule add_mono)
+        apply (rule d(2)[of "x" "u",unfolded o_def])
+        prefer 4
+        apply (rule d(2)[of "x" "v",unfolded o_def])
+        using ball[rule_format,of u] ball[rule_format,of v]
+        using xk(1-2)
+        unfolding k subset_eq
+        apply (auto simp add:dist_real_def)
+        done
+      also have "\<dots> \<le> e * (Sup k - Inf k)"
+        unfolding k interval_bounds_real[OF *]
+        using xk(1)
+        unfolding k
+        by (auto simp add: dist_real_def field_simps)
+      finally show "norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le>
+        e * (Sup k - Inf k)"
+        unfolding box_real k interval_bounds_real[OF *] content_real[OF *]
+          interval_upperbound_real interval_lowerbound_real
+          .
+    qed
+  qed
+qed
+
+lemma ident_has_integral:
+  fixes a::real
+  assumes "a \<le> b"
+  shows "((\<lambda>x. x) has_integral (b\<^sup>2 - a\<^sup>2) / 2) {a..b}"
+proof -
+  have "((\<lambda>x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}"
+    apply (rule fundamental_theorem_of_calculus [OF assms], clarify)
+    unfolding power2_eq_square
+    by (rule derivative_eq_intros | simp)+
+  then show ?thesis
+    by (simp add: field_simps)
+qed
+
+lemma integral_ident [simp]:
+  fixes a::real
+  assumes "a \<le> b"
+  shows "integral {a..b} (\<lambda>x. x) = (if a \<le> b then (b\<^sup>2 - a\<^sup>2) / 2 else 0)"
+using ident_has_integral integral_unique by fastforce
+
+lemma ident_integrable_on:
+  fixes a::real
+  shows "(\<lambda>x. x) integrable_on {a..b}"
+by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral)
+
+
+subsection \<open>Taylor series expansion\<close>
+
+lemma (in bounded_bilinear) setsum_prod_derivatives_has_vector_derivative:
+  assumes "p>0"
+  and f0: "Df 0 = f"
+  and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
+    (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})"
+  and g0: "Dg 0 = g"
+  and Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
+    (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})"
+  and ivl: "a \<le> t" "t \<le> b"
+  shows "((\<lambda>t. \<Sum>i<p. (-1)^i *\<^sub>R prod (Df i t) (Dg (p - Suc i) t))
+    has_vector_derivative
+      prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t))
+    (at t within {a .. b})"
+  using assms
+proof cases
+  assume p: "p \<noteq> 1"
+  define p' where "p' = p - 2"
+  from assms p have p': "{..<p} = {..Suc p'}" "p = Suc (Suc p')"
+    by (auto simp: p'_def)
+  have *: "\<And>i. i \<le> p' \<Longrightarrow> Suc (Suc p' - i) = (Suc (Suc p') - i)"
+    by auto
+  let ?f = "\<lambda>i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))"
+  have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
+    prod (Df (Suc i) t) (Dg (p - Suc i) t))) =
+    (\<Sum>i\<le>(Suc p'). ?f i - ?f (Suc i))"
+    by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost)
+  also note setsum_telescope
+  finally
+  have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
+    prod (Df (Suc i) t) (Dg (p - Suc i) t)))
+    = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)"
+    unfolding p'[symmetric]
+    by (simp add: assms)
+  thus ?thesis
+    using assms
+    by (auto intro!: derivative_eq_intros has_vector_derivative)
+qed (auto intro!: derivative_eq_intros has_vector_derivative)
+
+lemma
+  fixes f::"real\<Rightarrow>'a::banach"
+  assumes "p>0"
+  and f0: "Df 0 = f"
+  and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
+    (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})"
+  and ivl: "a \<le> b"
+  defines "i \<equiv> \<lambda>x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x"
+  shows taylor_has_integral:
+    "(i has_integral f b - (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a)) {a..b}"
+  and taylor_integral:
+    "f b = (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a) + integral {a..b} i"
+  and taylor_integrable:
+    "i integrable_on {a .. b}"
+proof goal_cases
+  case 1
+  interpret bounded_bilinear "scaleR::real\<Rightarrow>'a\<Rightarrow>'a"
+    by (rule bounded_bilinear_scaleR)
+  define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s
+  define Dg where [abs_def]:
+    "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s
+  have g0: "Dg 0 = g"
+    using \<open>p > 0\<close>
+    by (auto simp add: Dg_def divide_simps g_def split: if_split_asm)
+  {
+    fix m
+    assume "p > Suc m"
+    hence "p - Suc m = Suc (p - Suc (Suc m))"
+      by auto
+    hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)"
+      by auto
+  } note fact_eq = this
+  have Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
+    (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})"
+    unfolding Dg_def
+    by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps)
+  let ?sum = "\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t"
+  from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
+      OF \<open>p > 0\<close> g0 Dg f0 Df]
+  have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
+    (?sum has_vector_derivative
+      g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})"
+    by auto
+  from fundamental_theorem_of_calculus[rule_format, OF \<open>a \<le> b\<close> deriv]
+  have "(i has_integral ?sum b - ?sum a) {a .. b}"
+    by (simp add: i_def g_def Dg_def)
+  also
+  have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)"
+    and "{..<p} \<inter> {i. p = Suc i} = {p - 1}"
+    for p'
+    using \<open>p > 0\<close>
+    by (auto simp: power_mult_distrib[symmetric])
+  then have "?sum b = f b"
+    using Suc_pred'[OF \<open>p > 0\<close>]
+    by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib
+        cond_application_beta setsum.If_cases f0)
+  also
+  have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
+  proof safe
+    fix x
+    assume "x < p"
+    thus "x \<in> (\<lambda>x. p - x - 1) ` {..<p}"
+      by (auto intro!: image_eqI[where x = "p - x - 1"])
+  qed simp
+  from _ this
+  have "?sum a = (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a)"
+    by (rule setsum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
+  finally show c: ?case .
+  case 2 show ?case using c integral_unique by force
+  case 3 show ?case using c by force
+qed
+
+
+subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
+
+lemma gauge_modify:
+  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
+  shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
+  using assms
+  unfolding gauge_def
+  apply safe
+  defer
+  apply (erule_tac x="f x" in allE)
+  apply (erule_tac x="d (f x)" in allE)
+  apply auto
+  done
+
+
+subsection \<open>Only need trivial subintervals if the interval itself is trivial.\<close>
+
+lemma division_of_nontrivial:
+  fixes s :: "'a::euclidean_space set set"
+  assumes "s division_of (cbox a b)"
+    and "content (cbox a b) \<noteq> 0"
+  shows "{k. k \<in> s \<and> content k \<noteq> 0} division_of (cbox a b)"
+  using assms(1)
+  apply -
+proof (induct "card s" arbitrary: s rule: nat_less_induct)
+  fix s::"'a set set"
+  assume assm: "s division_of (cbox a b)"
+    "\<forall>m<card s. \<forall>x. m = card x \<longrightarrow>
+      x division_of (cbox a b) \<longrightarrow> {k \<in> x. content k \<noteq> 0} division_of (cbox a b)"
+  note s = division_ofD[OF assm(1)]
+  let ?thesis = "{k \<in> s. content k \<noteq> 0} division_of (cbox a b)"
+  {
+    presume *: "{k \<in> s. content k \<noteq> 0} \<noteq> s \<Longrightarrow> ?thesis"
+    show ?thesis
+      apply cases
+      defer
+      apply (rule *)
+      apply assumption
+      using assm(1)
+      apply auto
+      done
+  }
+  assume noteq: "{k \<in> s. content k \<noteq> 0} \<noteq> s"
+  then obtain k where k: "k \<in> s" "content k = 0"
+    by auto
+  from s(4)[OF k(1)] guess c d by (elim exE) note k=k this
+  from k have "card s > 0"
+    unfolding card_gt_0_iff using assm(1) by auto
+  then have card: "card (s - {k}) < card s"
+    using assm(1) k(1)
+    apply (subst card_Diff_singleton_if)
+    apply auto
+    done
+  have *: "closed (\<Union>(s - {k}))"
+    apply (rule closed_Union)
+    defer
+    apply rule
+    apply (drule DiffD1,drule s(4))
+    using assm(1)
+    apply auto
+    done
+  have "k \<subseteq> \<Union>(s - {k})"
+    apply safe
+    apply (rule *[unfolded closed_limpt,rule_format])
+    unfolding islimpt_approachable
+  proof safe
+    fix x
+    fix e :: real
+    assume as: "x \<in> k" "e > 0"
+    from k(2)[unfolded k content_eq_0] guess i ..
+    then have i:"c\<bullet>i = d\<bullet>i" "i\<in>Basis"
+      using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto
+    then have xi: "x\<bullet>i = d\<bullet>i"
+      using as unfolding k mem_box by (metis antisym)
+    define y where "y = (\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i) / 2 then c\<bullet>i +
+      min e (b\<bullet>i - c\<bullet>i) / 2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i) / 2 else x\<bullet>j) *\<^sub>R j)"
+    show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e"
+      apply (rule_tac x=y in bexI)
+    proof
+      have "d \<in> cbox c d"
+        using s(3)[OF k(1)]
+        unfolding k box_eq_empty mem_box
+        by (fastforce simp add: not_less)
+      then have "d \<in> cbox a b"
+        using s(2)[OF k(1)]
+        unfolding k
+        by auto
+      note di = this[unfolded mem_box,THEN bspec[where x=i]]
+      then have xyi: "y\<bullet>i \<noteq> x\<bullet>i"
+        unfolding y_def i xi
+        using as(2) assms(2)[unfolded content_eq_0] i(2)
+        by (auto elim!: ballE[of _ _ i])
+      then show "y \<noteq> x"
+        unfolding euclidean_eq_iff[where 'a='a] using i by auto
+      have *: "Basis = insert i (Basis - {i})"
+        using i by auto
+      have "norm (y - x) < e + setsum (\<lambda>i. 0) Basis"
+        apply (rule le_less_trans[OF norm_le_l1])
+        apply (subst *)
+        apply (subst setsum.insert)
+        prefer 3
+        apply (rule add_less_le_mono)
+      proof -
+        show "\<bar>(y - x) \<bullet> i\<bar> < e"
+          using di as(2) y_def i xi by (auto simp: inner_simps)
+        show "(\<Sum>i\<in>Basis - {i}. \<bar>(y - x) \<bullet> i\<bar>) \<le> (\<Sum>i\<in>Basis. 0)"
+          unfolding y_def by (auto simp: inner_simps)
+      qed auto
+      then show "dist y x < e"
+        unfolding dist_norm by auto
+      have "y \<notin> k"
+        unfolding k mem_box
+        apply rule
+        apply (erule_tac x=i in ballE)
+        using xyi k i xi
+        apply auto
+        done
+      moreover
+      have "y \<in> \<Union>s"
+        using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i
+        unfolding s mem_box y_def
+        by (auto simp: field_simps elim!: ballE[of _ _ i])
+      ultimately
+      show "y \<in> \<Union>(s - {k})" by auto
+    qed
+  qed
+  then have "\<Union>(s - {k}) = cbox a b"
+    unfolding s(6)[symmetric] by auto
+  then have  "{ka \<in> s - {k}. content ka \<noteq> 0} division_of (cbox a b)"
+    apply -
+    apply (rule assm(2)[rule_format,OF card refl])
+    apply (rule division_ofI)
+    defer
+    apply (rule_tac[1-4] s)
+    using assm(1)
+    apply auto
+    done
+  moreover
+  have "{ka \<in> s - {k}. content ka \<noteq> 0} = {k \<in> s. content k \<noteq> 0}"
+    using k by auto
+  ultimately show ?thesis by auto
+qed
+
+
+subsection \<open>Integrability on subintervals.\<close>
+
+lemma operative_integrable:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  shows "comm_monoid.operative op \<and> True (\<lambda>i. f integrable_on i)"
+  unfolding comm_monoid.operative_def[OF comm_monoid_and]
+  apply safe
+  apply (subst integrable_on_def)
+  unfolding has_integral_null_eq
+  apply (rule, rule refl)
+  apply (rule, assumption, assumption)+
+  unfolding integrable_on_def
+  by (auto intro!: has_integral_split)
+
+lemma integrable_subinterval:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on cbox a b"
+    and "cbox c d \<subseteq> cbox a b"
+  shows "f integrable_on cbox c d"
+  apply (cases "cbox c d = {}")
+  defer
+  apply (rule partial_division_extend_1[OF assms(2)],assumption)
+  using comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable,symmetric,of _ _ _ f] assms(1)
+  apply (auto simp: comm_monoid_set_F_and)
+  done
+
+lemma integrable_subinterval_real:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "f integrable_on {a .. b}"
+    and "{c .. d} \<subseteq> {a .. b}"
+  shows "f integrable_on {c .. d}"
+  by (metis assms(1) assms(2) box_real(2) integrable_subinterval)
+
+
+subsection \<open>Combining adjacent intervals in 1 dimension.\<close>
+
+lemma has_integral_combine:
+  fixes a b c :: real
+  assumes "a \<le> c"
+    and "c \<le> b"
+    and "(f has_integral i) {a .. c}"
+    and "(f has_integral (j::'a::banach)) {c .. b}"
+  shows "(f has_integral (i + j)) {a .. b}"
+proof -
+  note operative_integral[of f, unfolded comm_monoid.operative_1_le[OF add.comm_monoid_lift_option]]
+  note conjunctD2[OF this,rule_format]
+  note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]]
+  then have "f integrable_on cbox a b"
+    apply -
+    apply (rule ccontr)
+    apply (subst(asm) if_P)
+    defer
+    apply (subst(asm) if_P)
+    using assms(3-)
+    apply auto
+    done
+  with *
+  show ?thesis
+    apply -
+    apply (subst(asm) if_P)
+    defer
+    apply (subst(asm) if_P)
+    defer
+    apply (subst(asm) if_P)
+    using assms(3-)
+    apply (auto simp add: integrable_on_def integral_unique)
+    done
+qed
+
+lemma integral_combine:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "a \<le> c"
+    and "c \<le> b"
+    and "f integrable_on {a .. b}"
+  shows "integral {a .. c} f + integral {c .. b} f = integral {a .. b} f"
+  apply (rule integral_unique[symmetric])
+  apply (rule has_integral_combine[OF assms(1-2)])
+  apply (metis assms(2) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel2 monoid_add_class.add.left_neutral)
+  by (metis assms(1) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel1 monoid_add_class.add.right_neutral)
+
+lemma integrable_combine:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "a \<le> c"
+    and "c \<le> b"
+    and "f integrable_on {a .. c}"
+    and "f integrable_on {c .. b}"
+  shows "f integrable_on {a .. b}"
+  using assms
+  unfolding integrable_on_def
+  by (fastforce intro!:has_integral_combine)
+
+
+subsection \<open>Reduce integrability to "local" integrability.\<close>
+
+lemma integrable_on_little_subintervals:
+  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "\<forall>x\<in>cbox a b. \<exists>d>0. \<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
+    f integrable_on cbox u v"
+  shows "f integrable_on cbox a b"
+proof -
+  have "\<forall>x. \<exists>d. x\<in>cbox a b \<longrightarrow> d>0 \<and> (\<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
+    f integrable_on cbox u v)"
+    using assms by auto
+  note this[unfolded gauge_existence_lemma]
+  from choice[OF this] guess d .. note d=this[rule_format]
+  guess p
+    apply (rule fine_division_exists[OF gauge_ball_dependent,of d a b])
+    using d
+    by auto
+  note p=this(1-2)
+  note division_of_tagged_division[OF this(1)]
+  note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable, OF this, symmetric, of f]
+  show ?thesis
+    unfolding * comm_monoid_set_F_and
+    apply safe
+    unfolding snd_conv
+  proof -
+    fix x k
+    assume "(x, k) \<in> p"
+    note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this]
+    then show "f integrable_on k"
+      apply safe
+      apply (rule d[THEN conjunct2,rule_format,of x])
+      apply (auto intro: order.trans)
+      done
+  qed
+qed
+
+
+subsection \<open>Second FTC or existence of antiderivative.\<close>
+
+lemma integrable_const[intro]: "(\<lambda>x. c) integrable_on cbox a b"
+  unfolding integrable_on_def
+  apply rule
+  apply (rule has_integral_const)
+  done
+
+lemma integral_has_vector_derivative_continuous_at:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes f: "f integrable_on {a..b}"
+      and x: "x \<in> {a..b}"
+      and fx: "continuous (at x within {a..b}) f"
+  shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
+proof -
+  let ?I = "\<lambda>a b. integral {a..b} f"
+  { fix e::real
+    assume "e > 0"
+    obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {a..b}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> norm(f x' - f x) \<le> e"
+      using \<open>e>0\<close> fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le)
+    have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
+           if y: "y \<in> {a..b}" and yx: "\<bar>y - x\<bar> < d" for y
+    proof (cases "y < x")
+      case False
+      have "f integrable_on {a..y}"
+        using f y by (simp add: integrable_subinterval_real)
+      then have Idiff: "?I a y - ?I a x = ?I x y"
+        using False x by (simp add: algebra_simps integral_combine)
+      have fux_int: "((\<lambda>u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}"
+        apply (rule has_integral_sub)
+        using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]])
+        using has_integral_const_real [of "f x" x y] False
+        apply (simp add: )
+        done
+      show ?thesis
+        using False
+        apply (simp add: abs_eq_content del: content_real_if)
+        apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
+        using yx False d x y \<open>e>0\<close> apply (auto simp add: Idiff fux_int)
+        done
+    next
+      case True
+      have "f integrable_on {a..x}"
+        using f x by (simp add: integrable_subinterval_real)
+      then have Idiff: "?I a x - ?I a y = ?I y x"
+        using True x y by (simp add: algebra_simps integral_combine)
+      have fux_int: "((\<lambda>u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}"
+        apply (rule has_integral_sub)
+        using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]])
+        using has_integral_const_real [of "f x" y x] True
+        apply (simp add: )
+        done
+      have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
+        using True
+        apply (simp add: abs_eq_content del: content_real_if)
+        apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
+        using yx True d x y \<open>e>0\<close> apply (auto simp add: Idiff fux_int)
+        done
+      then show ?thesis
+        by (simp add: algebra_simps norm_minus_commute)
+    qed
+    then have "\<exists>d>0. \<forall>y\<in>{a..b}. \<bar>y - x\<bar> < d \<longrightarrow> norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
+      using \<open>d>0\<close> by blast
+  }
+  then show ?thesis
+    by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left)
+qed
+
+lemma integral_has_vector_derivative:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a .. b} f"
+    and "x \<in> {a .. b}"
+  shows "((\<lambda>u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})"
+apply (rule integral_has_vector_derivative_continuous_at [OF integrable_continuous_real])
+using assms
+apply (auto simp: continuous_on_eq_continuous_within)
+done
+
+lemma antiderivative_continuous:
+  fixes q b :: real
+  assumes "continuous_on {a .. b} f"
+  obtains g where "\<forall>x\<in>{a .. b}. (g has_vector_derivative (f x::_::banach)) (at x within {a .. b})"
+  apply (rule that)
+  apply rule
+  using integral_has_vector_derivative[OF assms]
+  apply auto
+  done
+
+
+subsection \<open>Combined fundamental theorem of calculus.\<close>
+
+lemma antiderivative_integral_continuous:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "continuous_on {a .. b} f"
+  obtains g where "\<forall>u\<in>{a .. b}. \<forall>v \<in> {a .. b}. u \<le> v \<longrightarrow> (f has_integral (g v - g u)) {u .. v}"
+proof -
+  from antiderivative_continuous[OF assms] guess g . note g=this
+  show ?thesis
+    apply (rule that[of g])
+    apply safe
+  proof goal_cases
+    case prems: (1 u v)
+    have "\<forall>x\<in>cbox u v. (g has_vector_derivative f x) (at x within cbox u v)"
+      apply rule
+      apply (rule has_vector_derivative_within_subset)
+      apply (rule g[rule_format])
+      using prems(1,2)
+      apply auto
+      done
+    then show ?case
+      using fundamental_theorem_of_calculus[OF prems(3), of g f] by auto
+  qed
+qed
+
+
+subsection \<open>General "twiddling" for interval-to-interval function image.\<close>
+
+lemma has_integral_twiddle:
+  assumes "0 < r"
+    and "\<forall>x. h(g x) = x"
+    and "\<forall>x. g(h x) = x"
+    and "\<forall>x. continuous (at x) g"
+    and "\<forall>u v. \<exists>w z. g ` cbox u v = cbox w z"
+    and "\<forall>u v. \<exists>w z. h ` cbox u v = cbox w z"
+    and "\<forall>u v. content(g ` cbox u v) = r * content (cbox u v)"
+    and "(f has_integral i) (cbox a b)"
+  shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)"
+proof -
+  show ?thesis when *: "cbox a b \<noteq> {} \<Longrightarrow> ?thesis"
+    apply cases
+    defer
+    apply (rule *)
+    apply assumption
+  proof goal_cases
+    case prems: 1
+    then show ?thesis
+      unfolding prems assms(8)[unfolded prems has_integral_empty_eq] by auto
+  qed
+  assume "cbox a b \<noteq> {}"
+  from assms(6)[rule_format,of a b] guess w z by (elim exE) note wz=this
+  have inj: "inj g" "inj h"
+    unfolding inj_on_def
+    apply safe
+    apply(rule_tac[!] ccontr)
+    using assms(2)
+    apply(erule_tac x=x in allE)
+    using assms(2)
+    apply(erule_tac x=y in allE)
+    defer
+    using assms(3)
+    apply (erule_tac x=x in allE)
+    using assms(3)
+    apply(erule_tac x=y in allE)
+    apply auto
+    done
+  show ?thesis
+    unfolding has_integral_def has_integral_compact_interval_def
+    apply (subst if_P)
+    apply rule
+    apply rule
+    apply (rule wz)
+  proof safe
+    fix e :: real
+    assume e: "e > 0"
+    with assms(1) have "e * r > 0" by simp
+    from assms(8)[unfolded has_integral,rule_format,OF this] guess d by (elim exE conjE) note d=this[rule_format]
+    define d' where "d' x = {y. g y \<in> d (g x)}" for x
+    have d': "\<And>x. d' x = {y. g y \<in> (d (g x))}"
+      unfolding d'_def ..
+    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)"
+    proof (rule_tac x=d' in exI, safe)
+      show "gauge d'"
+        using d(1)
+        unfolding gauge_def d'
+        using continuous_open_preimage_univ[OF assms(4)]
+        by auto
+      fix p
+      assume as: "p tagged_division_of h ` cbox a b" "d' fine p"
+      note p = tagged_division_ofD[OF as(1)]
+      have "(\<lambda>(x, k). (g x, g ` k)) ` p tagged_division_of (cbox a b) \<and> d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
+        unfolding tagged_division_of
+      proof safe
+        show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)"
+          using as by auto
+        show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
+          using as(2) unfolding fine_def d' by auto
+        fix x k
+        assume xk[intro]: "(x, k) \<in> p"
+        show "g x \<in> g ` k"
+          using p(2)[OF xk] by auto
+        show "\<exists>u v. g ` k = cbox u v"
+          using p(4)[OF xk] using assms(5-6) by auto
+        {
+          fix y
+          assume "y \<in> k"
+          then show "g y \<in> cbox a b" "g y \<in> cbox a b"
+            using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"]
+            using assms(2)[rule_format,of y]
+            unfolding inj_image_mem_iff[OF inj(2)]
+            by auto
+        }
+        fix x' k'
+        assume xk': "(x', k') \<in> p"
+        fix z
+        assume z: "z \<in> interior (g ` k)" "z \<in> interior (g ` k')"
+        have same: "(x, k) = (x', k')"
+          apply -
+          apply (rule ccontr)
+          apply (drule p(5)[OF xk xk'])
+        proof -
+          assume as: "interior k \<inter> interior k' = {}"
+          have "z \<in> g ` (interior k \<inter> interior k')"
+            using interior_image_subset[OF assms(4) inj(1)] z
+            unfolding image_Int[OF inj(1)] by blast
+          then show False
+            using as by blast
+        qed
+        then show "g x = g x'"
+          by auto
+        {
+          fix z
+          assume "z \<in> k"
+          then show "g z \<in> g ` k'"
+            using same by auto
+        }
+        {
+          fix z
+          assume "z \<in> k'"
+          then show "g z \<in> g ` k"
+            using same by auto
+        }
+      next
+        fix x
+        assume "x \<in> cbox a b"
+        then have "h x \<in>  \<Union>{k. \<exists>x. (x, k) \<in> p}"
+          using p(6) by auto
+        then guess X unfolding Union_iff .. note X=this
+        from this(1) guess y unfolding mem_Collect_eq ..
+        then show "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}"
+          apply -
+          apply (rule_tac X="g ` X" in UnionI)
+          defer
+          apply (rule_tac x="h x" in image_eqI)
+          using X(2) assms(3)[rule_format,of x]
+          apply auto
+          done
+      qed
+        note ** = d(2)[OF this]
+        have *: "inj_on (\<lambda>(x, k). (g x, g ` k)) p"
+          using inj(1) unfolding inj_on_def by fastforce
+        have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _")
+          using assms(7)
+          apply (simp only: algebra_simps add_left_cancel scaleR_right.setsum)
+          apply (subst setsum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
+          apply (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4))
+          done
+      also have "\<dots> = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r")
+        unfolding scaleR_diff_right scaleR_scaleR
+        using assms(1)
+        by auto
+      finally have *: "?l = ?r" .
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e"
+        using **
+        unfolding *
+        unfolding norm_scaleR
+        using assms(1)
+        by (auto simp add:field_simps)
+    qed
+  qed
+qed
+
+
+subsection \<open>Special case of a basic affine transformation.\<close>
+
+lemma interval_image_affinity_interval:
+  "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v"
+  unfolding image_affinity_cbox
+  by auto
+
+lemma content_image_affinity_cbox:
+  "content((\<lambda>x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) =
+    \<bar>m\<bar> ^ DIM('a) * content (cbox a b)" (is "?l = ?r")
+proof (cases "cbox a b = {}")
+  case True then show ?thesis by simp
+next
+  case False
+  show ?thesis
+  proof (cases "m \<ge> 0")
+    case True
+    with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \<noteq> {}"
+      unfolding box_ne_empty
+      apply (intro ballI)
+      apply (erule_tac x=i in ballE)
+      apply (auto simp: inner_simps mult_left_mono)
+      done
+    moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b - a) \<bullet> i"
+      by (simp add: inner_simps field_simps)
+    ultimately show ?thesis
+      by (simp add: image_affinity_cbox True content_cbox'
+        setprod.distrib setprod_constant inner_diff_left)
+  next
+    case False
+    with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
+      unfolding box_ne_empty
+      apply (intro ballI)
+      apply (erule_tac x=i in ballE)
+      apply (auto simp: inner_simps mult_left_mono)
+      done
+    moreover from False have *: "\<And>i. (m *\<^sub>R a + c) \<bullet> i - (m *\<^sub>R b + c) \<bullet> i = (-m) *\<^sub>R (b - a) \<bullet> i"
+      by (simp add: inner_simps field_simps)
+    ultimately show ?thesis using False
+      by (simp add: image_affinity_cbox content_cbox'
+        setprod.distrib[symmetric] setprod_constant[symmetric] inner_diff_left)
+  qed
+qed
+
+lemma has_integral_affinity:
+  fixes a :: "'a::euclidean_space"
+  assumes "(f has_integral i) (cbox a b)"
+      and "m \<noteq> 0"
+  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (\<bar>m\<bar> ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)"
+  apply (rule has_integral_twiddle)
+  using assms
+  apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox)
+  apply (rule zero_less_power)
+  unfolding scaleR_right_distrib
+  apply auto
+  done
+
+lemma integrable_affinity:
+  assumes "f integrable_on cbox a b"
+    and "m \<noteq> 0"
+  shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)"
+  using assms
+  unfolding integrable_on_def
+  apply safe
+  apply (drule has_integral_affinity)
+  apply auto
+  done
+
+lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified]
+
+subsection \<open>Special case of stretching coordinate axes separately.\<close>
+
+lemma content_image_stretch_interval:
+  "content ((\<lambda>x::'a::euclidean_space. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)::'a) ` cbox a b) =
+    \<bar>setprod m Basis\<bar> * content (cbox a b)"
+proof (cases "cbox a b = {}")
+  case True
+  then show ?thesis
+    unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
+next
+  case False
+  then have "(\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b \<noteq> {}"
+    by auto
+  then show ?thesis
+    using False
+    unfolding content_def image_stretch_interval
+    apply -
+    unfolding interval_bounds' if_not_P
+    unfolding abs_setprod setprod.distrib[symmetric]
+    apply (rule setprod.cong)
+    apply (rule refl)
+    unfolding lessThan_iff
+    apply (simp only: inner_setsum_left_Basis)
+  proof -
+    fix i :: 'a
+    assume i: "i \<in> Basis"
+    have "(m i < 0 \<or> m i > 0) \<or> m i = 0"
+      by auto
+    then show "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) - min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) =
+      \<bar>m i\<bar> * (b \<bullet> i - a \<bullet> i)"
+      apply -
+      apply (erule disjE)+
+      unfolding min_def max_def
+      using False[unfolded box_ne_empty,rule_format,of i] i
+      apply (auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos)
+      done
+  qed
+qed
+
+lemma has_integral_stretch:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "(f has_integral i) (cbox a b)"
+    and "\<forall>k\<in>Basis. m k \<noteq> 0"
+  shows "((\<lambda>x. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) has_integral
+    ((1/ \<bar>setprod m Basis\<bar>) *\<^sub>R i)) ((\<lambda>x. (\<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b)"
+  apply (rule has_integral_twiddle[where f=f])
+  unfolding zero_less_abs_iff content_image_stretch_interval
+  unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a]
+  using assms
+proof -
+  show "\<forall>y::'a. continuous (at y) (\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k))"
+    apply rule
+    apply (rule linear_continuous_at)
+    unfolding linear_linear
+    unfolding linear_iff inner_simps euclidean_eq_iff[where 'a='a]
+    apply (auto simp add: field_simps)
+    done
+qed auto
+
+lemma integrable_stretch:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
+  assumes "f integrable_on cbox a b"
+    and "\<forall>k\<in>Basis. m k \<noteq> 0"
+  shows "(\<lambda>x::'a. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) integrable_on
+    ((\<lambda>x. \<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k) ` cbox a b)"
+  using assms
+  unfolding integrable_on_def
+  apply -
+  apply (erule exE)
+  apply (drule has_integral_stretch)
+  apply assumption
+  apply auto
+  done
+
+subsection \<open>even more special cases.\<close>
+
+lemma uminus_interval_vector[simp]:
+  fixes a b :: "'a::euclidean_space"
+  shows "uminus ` cbox a b = cbox (-b) (-a)"
+  apply (rule set_eqI)
+  apply rule
+  defer
+  unfolding image_iff
+  apply (rule_tac x="-x" in bexI)
+  apply (auto simp add:minus_le_iff le_minus_iff mem_box)
+  done
+
+lemma has_integral_reflect_lemma[intro]:
+  assumes "(f has_integral i) (cbox a b)"
+  shows "((\<lambda>x. f(-x)) has_integral i) (cbox (-b) (-a))"
+  using has_integral_affinity[OF assms, of "-1" 0]
+  by auto
+
+lemma has_integral_reflect_lemma_real[intro]:
+  assumes "(f has_integral i) {a .. b::real}"
+  shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
+  using assms
+  unfolding box_real[symmetric]
+  by (rule has_integral_reflect_lemma)
+
+lemma has_integral_reflect[simp]:
+  "((\<lambda>x. f (-x)) has_integral i) (cbox (-b) (-a)) \<longleftrightarrow> (f has_integral i) (cbox a b)"
+  apply rule
+  apply (drule_tac[!] has_integral_reflect_lemma)
+  apply auto
+  done
+
+lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on cbox (-b) (-a) \<longleftrightarrow> f integrable_on cbox a b"
+  unfolding integrable_on_def by auto
+
+lemma integrable_reflect_real[simp]: "(\<lambda>x. f(-x)) integrable_on {-b .. -a} \<longleftrightarrow> f integrable_on {a .. b::real}"
+  unfolding box_real[symmetric]
+  by (rule integrable_reflect)
+
+lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\<lambda>x. f (-x)) = integral (cbox a b) f"
+  unfolding integral_def by auto
+
+lemma integral_reflect_real[simp]: "integral {-b .. -a} (\<lambda>x. f (-x)) = integral {a .. b::real} f"
+  unfolding box_real[symmetric]
+  by (rule integral_reflect)
+
+
+subsection \<open>Stronger form of FCT; quite a tedious proof.\<close>
+
+lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
+  by (meson zero_less_one)
+
+lemma additive_tagged_division_1':
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b"
+    and "p tagged_division_of {a..b}"
+  shows "setsum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
+  using additive_tagged_division_1[OF _ assms(2), of f]
+  using assms(1)
+  by auto
+
+lemma split_minus[simp]: "(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
+  by (simp add: split_def)
+
+lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
+  apply (subst(asm)(2) norm_minus_cancel[symmetric])
+  apply (drule norm_triangle_le)
+  apply (auto simp add: algebra_simps)
+  done
+
+lemma fundamental_theorem_of_calculus_interior:
+  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
+  assumes "a \<le> b"
+    and "continuous_on {a .. b} f"
+    and "\<forall>x\<in>{a <..< b}. (f has_vector_derivative f'(x)) (at x)"
+  shows "(f' has_integral (f b - f a)) {a .. b}"
+proof -
+  {
+    presume *: "a < b \<Longrightarrow> ?thesis"
+    show ?thesis
+    proof (cases "a < b")
+      case True
+      then show ?thesis by (rule *)
+    next
+      case False
+      then have "a = b"
+        using assms(1) by auto
+      then have *: "cbox a b = {b}" "f b - f a = 0"
+        by (auto simp add:  order_antisym)
+      show ?thesis
+        unfolding *(2)
+        unfolding content_eq_0
+        using * \<open>a = b\<close>
+        by (auto simp: ex_in_conv)
+    qed
+  }
+  assume ab: "a < b"
+  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
+    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a .. b})"
+  { presume "\<And>e. e > 0 \<Longrightarrow> ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto }
+  fix e :: real
+  assume e: "e > 0"
+  note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib]
+  note conjunctD2[OF this]
+  note bounded=this(1) and this(2)
+  from this(2) have "\<forall>x\<in>box a b. \<exists>d>0. \<forall>y. norm (y - x) < d \<longrightarrow>
+    norm (f y - f x - (y - x) *\<^sub>R f' x) \<le> e/2 * norm (y - x)"
+    apply -
+    apply safe
+    apply (erule_tac x=x in ballE)
+    apply (erule_tac x="e/2" in allE)
+    using e
+    apply auto
+    done
+  note this[unfolded bgauge_existence_lemma]
+  from choice[OF this] guess d ..
+  note conjunctD2[OF this[rule_format]]
+  note d = this[rule_format]
+  have "bounded (f ` cbox a b)"
+    apply (rule compact_imp_bounded compact_continuous_image)+
+    using compact_cbox assms
+    apply auto
+    done
+  from this[unfolded bounded_pos] guess B .. note B = this[rule_format]
+
+  have "\<exists>da. 0 < da \<and> (\<forall>c. a \<le> c \<and> {a .. c} \<subseteq> {a .. b} \<and> {a .. c} \<subseteq> ball a da \<longrightarrow>
+    norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
+  proof -
+    have "a \<in> {a .. b}"
+      using ab by auto
+    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
+    note * = this[unfolded continuous_within Lim_within,rule_format]
+    have "(e * (b - a)) / 8 > 0"
+      using e ab by (auto simp add: field_simps)
+    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
+    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' a) \<le> (e * (b - a)) / 8"
+    proof (cases "f' a = 0")
+      case True
+      thus ?thesis using ab e by auto
+    next
+      case False
+      then show ?thesis
+        apply (rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI)
+        using ab e
+        apply (auto simp add: field_simps)
+        done
+    qed
+    then guess l .. note l = conjunctD2[OF this]
+    show ?thesis
+      apply (rule_tac x="min k l" in exI)
+      apply safe
+      unfolding min_less_iff_conj
+      apply rule
+      apply (rule l k)+
+    proof -
+      fix c
+      assume as: "a \<le> c" "{a .. c} \<subseteq> {a .. b}" "{a .. c} \<subseteq> ball a (min k l)"
+      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box]
+      have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)"
+        by (rule norm_triangle_ineq4)
+      also have "\<dots> \<le> e * (b - a) / 8 + e * (b - a) / 8"
+      proof (rule add_mono)
+        have "\<bar>c - a\<bar> \<le> \<bar>l\<bar>"
+          using as' by auto
+        then show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b - a) / 8"
+          apply -
+          apply (rule order_trans[OF _ l(2)])
+          unfolding norm_scaleR
+          apply (rule mult_right_mono)
+          apply auto
+          done
+      next
+        show "norm (f c - f a) \<le> e * (b - a) / 8"
+          apply (rule less_imp_le)
+          apply (cases "a = c")
+          defer
+          apply (rule k(2)[unfolded dist_norm])
+          using as' e ab
+          apply (auto simp add: field_simps)
+          done
+      qed
+      finally show "norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4"
+        unfolding content_real[OF as(1)] by auto
+    qed
+  qed
+  then guess da .. note da=conjunctD2[OF this,rule_format]
+
+  have "\<exists>db>0. \<forall>c\<le>b. {c .. b} \<subseteq> {a .. b} \<and> {c .. b} \<subseteq> ball b db \<longrightarrow>
+    norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
+  proof -
+    have "b \<in> {a .. b}"
+      using ab by auto
+    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
+    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0"
+      using e ab by (auto simp add: field_simps)
+    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
+    have "\<exists>l. 0 < l \<and> norm (l *\<^sub>R f' b) \<le> (e * (b - a)) / 8"
+    proof (cases "f' b = 0")
+      case True
+      thus ?thesis using ab e by auto
+    next
+      case False
+      then show ?thesis
+        apply (rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI)
+        using ab e
+        apply (auto simp add: field_simps)
+        done
+    qed
+    then guess l .. note l = conjunctD2[OF this]
+    show ?thesis
+      apply (rule_tac x="min k l" in exI)
+      apply safe
+      unfolding min_less_iff_conj
+      apply rule
+      apply (rule l k)+
+    proof -
+      fix c
+      assume as: "c \<le> b" "{c..b} \<subseteq> {a..b}" "{c..b} \<subseteq> ball b (min k l)"
+      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box]
+      have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)"
+        by (rule norm_triangle_ineq4)
+      also have "\<dots> \<le> e * (b - a) / 8 + e * (b - a) / 8"
+      proof (rule add_mono)
+        have "\<bar>c - b\<bar> \<le> \<bar>l\<bar>"
+          using as' by auto
+        then show "norm ((b - c) *\<^sub>R f' b) \<le> e * (b - a) / 8"
+          apply -
+          apply (rule order_trans[OF _ l(2)])
+          unfolding norm_scaleR
+          apply (rule mult_right_mono)
+          apply auto
+          done
+      next
+        show "norm (f b - f c) \<le> e * (b - a) / 8"
+          apply (rule less_imp_le)
+          apply (cases "b = c")
+          defer
+          apply (subst norm_minus_commute)
+          apply (rule k(2)[unfolded dist_norm])
+          using as' e ab
+          apply (auto simp add: field_simps)
+          done
+      qed
+      finally show "norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4"
+        unfolding content_real[OF as(1)] by auto
+    qed
+  qed
+  then guess db .. note db=conjunctD2[OF this,rule_format]
+
+  let ?d = "(\<lambda>x. ball x (if x=a then da else if x=b then db else d x))"
+  show "?P e"
+    apply (rule_tac x="?d" in exI)
+  proof (safe, goal_cases)
+    case 1
+    show ?case
+      apply (rule gauge_ball_dependent)
+      using ab db(1) da(1) d(1)
+      apply auto
+      done
+  next
+    case as: (2 p)
+    let ?A = "{t. fst t \<in> {a, b}}"
+    note p = tagged_division_ofD[OF as(1)]
+    have pA: "p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"
+      using as by auto
+    note * = additive_tagged_division_1'[OF assms(1) as(1), symmetric]
+    have **: "\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2"
+      by arith
+    show ?case
+      unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[symmetric] split_minus
+      unfolding setsum_right_distrib
+      apply (subst(2) pA)
+      apply (subst pA)
+      unfolding setsum.union_disjoint[OF pA(2-)]
+    proof (rule norm_triangle_le, rule **, goal_cases)
+      case 1
+      show ?case
+        apply (rule order_trans)
+        apply (rule setsum_norm_le)
+        defer
+        apply (subst setsum_divide_distrib)
+        apply (rule order_refl)
+        apply safe
+        apply (unfold not_le o_def split_conv fst_conv)
+      proof (rule ccontr)
+        fix x k
+        assume xk: "(x, k) \<in> p"
+          "e * (Sup k -  Inf k) / 2 <
+            norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))"
+        from p(4)[OF this(1)] guess u v by (elim exE) note k=this
+        then have "u \<le> v" and uv: "{u, v} \<subseteq> cbox u v"
+          using p(2)[OF xk(1)] by auto
+        note result = xk(2)[unfolded k box_real interval_bounds_real[OF this(1)] content_real[OF this(1)]]
+
+        assume as': "x \<noteq> a" "x \<noteq> b"
+        then have "x \<in> box a b"
+          using p(2-3)[OF xk(1)] by (auto simp: mem_box)
+        note  * = d(2)[OF this]
+        have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) =
+          norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))"
+          apply (rule arg_cong[of _ _ norm])
+          unfolding scaleR_left.diff
+          apply auto
+          done
+        also have "\<dots> \<le> e / 2 * norm (u - x) + e / 2 * norm (v - x)"
+          apply (rule norm_triangle_le_sub)
+          apply (rule add_mono)
+          apply (rule_tac[!] *)
+          using fineD[OF as(2) xk(1)] as'
+          unfolding k subset_eq
+          apply -
+          apply (erule_tac x=u in ballE)
+          apply (erule_tac[3] x=v in ballE)
+          using uv
+          apply (auto simp:dist_real_def)
+          done
+        also have "\<dots> \<le> e / 2 * norm (v - u)"
+          using p(2)[OF xk(1)]
+          unfolding k
+          by (auto simp add: field_simps)
+        finally have "e * (v - u) / 2 < e * (v - u) / 2"
+          apply -
+          apply (rule less_le_trans[OF result])
+          using uv
+          apply auto
+          done
+        then show False by auto
+      qed
+    next
+      have *: "\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2) / 2 \<Longrightarrow> x - s1 \<le> s2 / 2"
+        by auto
+      case 2
+      show ?case
+        apply (rule *)
+        apply (rule setsum_nonneg)
+        apply rule
+        apply (unfold split_paired_all split_conv)
+        defer
+        unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
+        unfolding setsum_right_distrib[symmetric]
+        apply (subst additive_tagged_division_1[OF _ as(1)])
+        apply (rule assms)
+      proof -
+        fix x k
+        assume "(x, k) \<in> p \<inter> {t. fst t \<in> {a, b}}"
+        note xk=IntD1[OF this]
+        from p(4)[OF this] guess u v by (elim exE) note uv=this
+        with p(2)[OF xk] have "cbox u v \<noteq> {}"
+          by auto
+        then show "0 \<le> e * ((Sup k) - (Inf k))"
+          unfolding uv using e by (auto simp add: field_simps)
+      next
+        have *: "\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm (setsum f t) \<le> e \<Longrightarrow> norm (setsum f s) \<le> e"
+          by auto
+        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x -
+          (f ((Sup k)) - f ((Inf k)))) \<le> e * (b - a) / 2"
+          apply (rule *[where t1="p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0}"])
+          apply (rule setsum.mono_neutral_right[OF pA(2)])
+          defer
+          apply rule
+          unfolding split_paired_all split_conv o_def
+        proof goal_cases
+          fix x k
+          assume "(x, k) \<in> p \<inter> {t. fst t \<in> {a, b}} - p \<inter> {t. fst t \<in> {a, b} \<and> content (snd t) \<noteq> 0}"
+          then have xk: "(x, k) \<in> p" "content k = 0"
+            by auto
+          from p(4)[OF xk(1)] guess u v by (elim exE) note uv=this
+          have "k \<noteq> {}"
+            using p(2)[OF xk(1)] by auto
+          then have *: "u = v"
+            using xk
+            unfolding uv content_eq_0 box_eq_empty
+            by auto
+          then show "content k *\<^sub>R (f' (x)) - (f ((Sup k)) - f ((Inf k))) = 0"
+            using xk unfolding uv by auto
+        next
+          have *: "p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0} =
+            {t. t\<in>p \<and> fst t = a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = b \<and> content(snd t) \<noteq> 0}"
+            by blast
+          have **: "norm (setsum f s) \<le> e"
+            if "\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y"
+            and "\<forall>x. x \<in> s \<longrightarrow> norm (f x) \<le> e"
+            and "e > 0"
+            for s f and e :: real
+          proof (cases "s = {}")
+            case True
+            with that show ?thesis by auto
+          next
+            case False
+            then obtain x where "x \<in> s"
+              by auto
+            then have *: "s = {x}"
+              using that(1) by auto
+            then show ?thesis
+              using \<open>x \<in> s\<close> that(2) by auto
+          qed
+          case 2
+          show ?case
+            apply (subst *)
+            apply (subst setsum.union_disjoint)
+            prefer 4
+            apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"])
+            apply (rule norm_triangle_le,rule add_mono)
+            apply (rule_tac[1-2] **)
+          proof -
+            let ?B = "\<lambda>x. {t \<in> p. fst t = x \<and> content (snd t) \<noteq> 0}"
+            have pa: "\<exists>v. k = cbox a v \<and> a \<le> v" if "(a, k) \<in> p" for k
+            proof -
+              guess u v using p(4)[OF that] by (elim exE) note uv=this
+              have *: "u \<le> v"
+                using p(2)[OF that] unfolding uv by auto
+              have u: "u = a"
+              proof (rule ccontr)
+                have "u \<in> cbox u v"
+                  using p(2-3)[OF that(1)] unfolding uv by auto
+                have "u \<ge> a"
+                  using p(2-3)[OF that(1)] unfolding uv subset_eq by auto
+                moreover assume "\<not> ?thesis"
+                ultimately have "u > a" by auto
+                then show False
+                  using p(2)[OF that(1)] unfolding uv by (auto simp add:)
+              qed
+              then show ?thesis
+                apply (rule_tac x=v in exI)
+                unfolding uv
+                using *
+                apply auto
+                done
+            qed
+            have pb: "\<exists>v. k = cbox v b \<and> b \<ge> v" if "(b, k) \<in> p" for k
+            proof -
+              guess u v using p(4)[OF that] by (elim exE) note uv=this
+              have *: "u \<le> v"
+                using p(2)[OF that] unfolding uv by auto
+              have u: "v = b"
+              proof (rule ccontr)
+                have "u \<in> cbox u v"
+                  using p(2-3)[OF that(1)] unfolding uv by auto
+                have "v \<le> b"
+                  using p(2-3)[OF that(1)] unfolding uv subset_eq by auto
+                moreover assume "\<not> ?thesis"
+                ultimately have "v < b" by auto
+                then show False
+                  using p(2)[OF that(1)] unfolding uv by (auto simp add:)
+              qed
+              then show ?thesis
+                apply (rule_tac x=u in exI)
+                unfolding uv
+                using *
+                apply auto
+                done
+            qed
+            show "\<forall>x y. x \<in> ?B a \<and> y \<in> ?B a \<longrightarrow> x = y"
+              apply (rule,rule,rule,unfold split_paired_all)
+              unfolding mem_Collect_eq fst_conv snd_conv
+              apply safe
+            proof -
+              fix x k k'
+              assume k: "(a, k) \<in> p" "(a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+              guess v using pa[OF k(1)] .. note v = conjunctD2[OF this]
+              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "min v v'"
+              have "box a ?v \<subseteq> k \<inter> k'"
+                unfolding v v' by (auto simp add: mem_box)
+              note interior_mono[OF this,unfolded interior_Int]
+              moreover have "(a + ?v)/2 \<in> box a ?v"
+                using k(3-)
+                unfolding v v' content_eq_0 not_le
+                by (auto simp add: mem_box)
+              ultimately have "(a + ?v)/2 \<in> interior k \<inter> interior k'"
+                unfolding interior_open[OF open_box] by auto
+              then have *: "k = k'"
+                apply -
+                apply (rule ccontr)
+                using p(5)[OF k(1-2)]
+                apply auto
+                done
+              { assume "x \<in> k" then show "x \<in> k'" unfolding * . }
+              { assume "x \<in> k'" then show "x \<in> k" unfolding * . }
+            qed
+            show "\<forall>x y. x \<in> ?B b \<and> y \<in> ?B b \<longrightarrow> x = y"
+              apply rule
+              apply rule
+              apply rule
+              apply (unfold split_paired_all)
+              unfolding mem_Collect_eq fst_conv snd_conv
+              apply safe
+            proof -
+              fix x k k'
+              assume k: "(b, k) \<in> p" "(b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
+              guess v using pb[OF k(1)] .. note v = conjunctD2[OF this]
+              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this]
+              let ?v = "max v v'"
+              have "box ?v b \<subseteq> k \<inter> k'"
+                unfolding v v' by (auto simp: mem_box)
+                note interior_mono[OF this,unfolded interior_Int]
+              moreover have " ((b + ?v)/2) \<in> box ?v b"
+                using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: mem_box)
+              ultimately have " ((b + ?v)/2) \<in> interior k \<inter> interior k'"
+                unfolding interior_open[OF open_box] by auto
+              then have *: "k = k'"
+                apply -
+                apply (rule ccontr)
+                using p(5)[OF k(1-2)]
+                apply auto
+                done
+              { assume "x \<in> k" then show "x \<in> k'" unfolding * . }
+              { assume "x \<in> k'" then show "x\<in>k" unfolding * . }
+            qed
+
+            let ?a = a and ?b = b (* a is something else while proofing the next theorem. *)
+            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' x - (f (Sup k) -
+              f (Inf k))) x) \<le> e * (b - a) / 4"
+              apply rule
+              apply rule
+              unfolding mem_Collect_eq
+              unfolding split_paired_all fst_conv snd_conv
+            proof (safe, goal_cases)
+              case prems: 1
+              guess v using pa[OF prems(1)] .. note v = conjunctD2[OF this]
+              have "?a \<in> {?a..v}"
+                using v(2) by auto
+              then have "v \<le> ?b"
+                using p(3)[OF prems(1)] unfolding subset_eq v by auto
+              moreover have "{?a..v} \<subseteq> ball ?a da"
+                using fineD[OF as(2) prems(1)]
+                apply -
+                apply (subst(asm) if_P)
+                apply (rule refl)
+                unfolding subset_eq
+                apply safe
+                apply (erule_tac x=" x" in ballE)
+                apply (auto simp add:subset_eq dist_real_def v)
+                done
+              ultimately show ?case
+                unfolding v interval_bounds_real[OF v(2)] box_real
+                apply -
+                apply(rule da(2)[of "v"])
+                using prems fineD[OF as(2) prems(1)]
+                unfolding v content_eq_0
+                apply auto
+                done
+            qed
+            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' x -
+              (f (Sup k) - f (Inf k))) x) \<le> e * (b - a) / 4"
+              apply rule
+              apply rule
+              unfolding mem_Collect_eq
+              unfolding split_paired_all fst_conv snd_conv
+            proof (safe, goal_cases)
+              case prems: 1
+              guess v using pb[OF prems(1)] .. note v = conjunctD2[OF this]
+              have "?b \<in> {v.. ?b}"
+                using v(2) by auto
+              then have "v \<ge> ?a" using p(3)[OF prems(1)]
+                unfolding subset_eq v by auto
+              moreover have "{v..?b} \<subseteq> ball ?b db"
+                using fineD[OF as(2) prems(1)]
+                apply -
+                apply (subst(asm) if_P, rule refl)
+                unfolding subset_eq
+                apply safe
+                apply (erule_tac x=" x" in ballE)
+                using ab
+                apply (auto simp add:subset_eq v dist_real_def)
+                done
+              ultimately show ?case
+                unfolding v
+                unfolding interval_bounds_real[OF v(2)] box_real
+                apply -
+                apply(rule db(2)[of "v"])
+                using prems fineD[OF as(2) prems(1)]
+                unfolding v content_eq_0
+                apply auto
+                done
+            qed
+          qed (insert p(1) ab e, auto simp add: field_simps)
+        qed auto
+      qed
+    qed
+  qed
+qed
+
+
+subsection \<open>Stronger form with finite number of exceptional points.\<close>
+
+lemma fundamental_theorem_of_calculus_interior_strong:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "finite s"
+    and "a \<le> b"
+    and "continuous_on {a .. b} f"
+    and "\<forall>x\<in>{a <..< b} - s. (f has_vector_derivative f'(x)) (at x)"
+  shows "(f' has_integral (f b - f a)) {a .. b}"
+  using assms
+proof (induct "card s" arbitrary: s a b)
+  case 0
+  show ?case
+    apply (rule fundamental_theorem_of_calculus_interior)
+    using 0
+    apply auto
+    done
+next
+  case (Suc n)
+  from this(2) guess c s'
+    apply -
+    apply (subst(asm) eq_commute)
+    unfolding card_Suc_eq
+    apply (subst(asm)(2) eq_commute)
+    apply (elim exE conjE)
+    done
+  note cs = this[rule_format]
+  show ?case
+  proof (cases "c \<in> box a b")
+    case False
+    then show ?thesis
+      apply -
+      apply (rule Suc(1)[OF cs(3) _ Suc(4,5)])
+      apply safe
+      defer
+      apply (rule Suc(6)[rule_format])
+      using Suc(3)
+      unfolding cs
+      apply auto
+      done
+  next
+    have *: "f b - f a = (f c - f a) + (f b - f c)"
+      by auto
+    case True
+    then have "a \<le> c" "c \<le> b"
+      by (auto simp: mem_box)
+    then show ?thesis
+      apply (subst *)
+      apply (rule has_integral_combine)
+      apply assumption+
+      apply (rule_tac[!] Suc(1)[OF cs(3)])
+      using Suc(3)
+      unfolding cs
+    proof -
+      show "continuous_on {a .. c} f" "continuous_on {c .. b} f"
+        apply (rule_tac[!] continuous_on_subset[OF Suc(5)])
+        using True
+        apply (auto simp: mem_box)
+        done
+      let ?P = "\<lambda>i j. \<forall>x\<in>{i <..< j} - s'. (f has_vector_derivative f' x) (at x)"
+      show "?P a c" "?P c b"
+        apply safe
+        apply (rule_tac[!] Suc(6)[rule_format])
+        using True
+        unfolding cs
+        apply (auto simp: mem_box)
+        done
+    qed auto
+  qed
+qed
+
+lemma fundamental_theorem_of_calculus_strong:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "finite s"
+    and "a \<le> b"
+    and "continuous_on {a .. b} f"
+    and "\<forall>x\<in>{a .. b} - s. (f has_vector_derivative f'(x)) (at x)"
+  shows "(f' has_integral (f b - f a)) {a .. b}"
+  apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f'])
+  using assms(4)
+  apply (auto simp: mem_box)
+  done
+
+lemma indefinite_integral_continuous_left:
+  fixes f:: "real \<Rightarrow> 'a::banach"
+  assumes "f integrable_on {a .. b}"
+    and "a < c"
+    and "c \<le> b"
+    and "e > 0"
+  obtains d where "d > 0"
+    and "\<forall>t. c - d < t \<and> t \<le> c \<longrightarrow> norm (integral {a .. c} f - integral {a .. t} f) < e"
+proof -
+  have "\<exists>w>0. \<forall>t. c - w < t \<and> t < c \<longrightarrow> norm (f c) * norm(c - t) < e / 3"
+  proof (cases "f c = 0")
+    case False
+    hence "0 < e / 3 / norm (f c)" using \<open>e>0\<close> by simp
+    then show ?thesis
+      apply -
+      apply rule
+      apply rule
+      apply assumption
+      apply safe
+    proof -
+      fix t
+      assume as: "t < c" and "c - e / 3 / norm (f c) < t"
+      then have "c - t < e / 3 / norm (f c)"
+        by auto
+      then have "norm (c - t) < e / 3 / norm (f c)"
+        using as by auto
+      then show "norm (f c) * norm (c - t) < e / 3"
+        using False
+        apply -
+        apply (subst mult.commute)
+        apply (subst pos_less_divide_eq[symmetric])
+        apply auto
+        done
+    qed
+  next
+    case True
+    show ?thesis
+      apply (rule_tac x=1 in exI)
+      unfolding True
+      using \<open>e > 0\<close>
+      apply auto
+      done
+  qed
+  then guess w .. note w = conjunctD2[OF this,rule_format]
+
+  have *: "e / 3 > 0"
+    using assms by auto
+  have "f integrable_on {a .. c}"
+    apply (rule integrable_subinterval_real[OF assms(1)])
+    using assms(2-3)
+    apply auto
+    done
+  from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 ..
+  note d1 = conjunctD2[OF this,rule_format]
+  define d where [abs_def]: "d x = ball x w \<inter> d1 x" for x
+  have "gauge d"
+    unfolding d_def using w(1) d1 by auto
+  note this[unfolded gauge_def,rule_format,of c]
+  note conjunctD2[OF this]
+  from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k ..
+  note k=conjunctD2[OF this]
+
+  let ?d = "min k (c - a) / 2"
+  show ?thesis
+    apply (rule that[of ?d])
+    apply safe
+  proof -
+    show "?d > 0"
+      using k(1) using assms(2) by auto
+    fix t
+    assume as: "c - ?d < t" "t \<le> c"
+    let ?thesis = "norm (integral ({a .. c}) f - integral ({a .. t}) f) < e"
+    {
+      presume *: "t < c \<Longrightarrow> ?thesis"
+      show ?thesis
+        apply (cases "t = c")
+        defer
+        apply (rule *)
+        apply (subst less_le)
+        using \<open>e > 0\<close> as(2)
+        apply auto
+        done
+    }
+    assume "t < c"
+
+    have "f integrable_on {a .. t}"
+      apply (rule integrable_subinterval_real[OF assms(1)])
+      using assms(2-3) as(2)
+      apply auto
+      done
+    from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 ..
+    note d2 = conjunctD2[OF this,rule_format]
+    define d3 where "d3 x = (if x \<le> t then d1 x \<inter> d2 x else d1 x)" for x
+    have "gauge d3"
+      using d2(1) d1(1) unfolding d3_def gauge_def by auto
+    from fine_division_exists_real[OF this, of a t] guess p . note p=this
+    note p'=tagged_division_ofD[OF this(1)]
+    have pt: "\<forall>(x,k)\<in>p. x \<le> t"
+    proof (safe, goal_cases)
+      case prems: 1
+      from p'(2,3)[OF prems] show ?case
+        by auto
+    qed
+    with p(2) have "d2 fine p"
+      unfolding fine_def d3_def
+      apply safe
+      apply (erule_tac x="(a,b)" in ballE)+
+      apply auto
+      done
+    note d2_fin = d2(2)[OF conjI[OF p(1) this]]
+
+    have *: "{a .. c} \<inter> {x. x \<bullet> 1 \<le> t} = {a .. t}" "{a .. c} \<inter> {x. x \<bullet> 1 \<ge> t} = {t .. c}"
+      using assms(2-3) as by (auto simp add: field_simps)
+    have "p \<union> {(c, {t .. c})} tagged_division_of {a .. c} \<and> d1 fine p \<union> {(c, {t .. c})}"
+      apply rule
+      apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"])
+      unfolding *
+      apply (rule p)
+      apply (rule tagged_division_of_self_real)
+      unfolding fine_def
+      apply safe
+    proof -
+      fix x k y
+      assume "(x,k) \<in> p" and "y \<in> k"
+      then show "y \<in> d1 x"
+        using p(2) pt
+        unfolding fine_def d3_def
+        apply -
+        apply (erule_tac x="(x,k)" in ballE)+
+        apply auto
+        done
+    next
+      fix x assume "x \<in> {t..c}"
+      then have "dist c x < k"
+        unfolding dist_real_def
+        using as(1)
+        by (auto simp add: field_simps)
+      then show "x \<in> d1 c"
+        using k(2)
+        unfolding d_def
+        by auto
+    qed (insert as(2), auto) note d1_fin = d1(2)[OF this]
+
+    have *: "integral {a .. c} f - integral {a .. t} f = -(((c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) -
+      integral {a .. c} f) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - integral {a .. t} f) + (c - t) *\<^sub>R f c"
+      "e = (e/3 + e/3) + e/3"
+      by auto
+    have **: "(\<Sum>(x, k)\<in>p \<union> {(c, {t .. c})}. content k *\<^sub>R f x) =
+      (c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
+    proof -
+      have **: "\<And>x F. F \<union> {x} = insert x F"
+        by auto
+      have "(c, cbox t c) \<notin> p"
+      proof (safe, goal_cases)
+        case prems: 1
+        from p'(2-3)[OF prems] have "c \<in> cbox a t"
+          by auto
+        then show False using \<open>t < c\<close>
+          by auto
+      qed
+      then show ?thesis
+        unfolding ** box_real
+        apply -
+        apply (subst setsum.insert)
+        apply (rule p')
+        unfolding split_conv
+        defer
+        apply (subst content_real)
+        using as(2)
+        apply auto
+        done
+    qed
+    have ***: "c - w < t \<and> t < c"
+    proof -
+      have "c - k < t"
+        using \<open>k>0\<close> as(1) by (auto simp add: field_simps)
+      moreover have "k \<le> w"
+        apply (rule ccontr)
+        using k(2)
+        unfolding subset_eq
+        apply (erule_tac x="c + ((k + w)/2)" in ballE)
+        unfolding d_def
+        using \<open>k > 0\<close> \<open>w > 0\<close>
+        apply (auto simp add: field_simps not_le not_less dist_real_def)
+        done
+      ultimately show ?thesis using \<open>t < c\<close>
+        by (auto simp add: field_simps)
+    qed
+    show ?thesis
+      unfolding *(1)
+      apply (subst *(2))
+      apply (rule norm_triangle_lt add_strict_mono)+
+      unfolding norm_minus_cancel
+      apply (rule d1_fin[unfolded **])
+      apply (rule d2_fin)
+      using w(2)[OF ***]
+      unfolding norm_scaleR
+      apply (auto simp add: field_simps)
+      done
+  qed
+qed
+
+lemma indefinite_integral_continuous_right:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "f integrable_on {a .. b}"
+    and "a \<le> c"
+    and "c < b"
+    and "e > 0"
+  obtains d where "0 < d"
+    and "\<forall>t. c \<le> t \<and> t < c + d \<longrightarrow> norm (integral {a .. c} f - integral {a .. t} f) < e"
+proof -
+  have *: "(\<lambda>x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \<le> - a"
+    using assms by auto
+  from indefinite_integral_continuous_left[OF * \<open>e>0\<close>] guess d . note d = this
+  let ?d = "min d (b - c)"
+  show ?thesis
+    apply (rule that[of "?d"])
+    apply safe
+  proof -
+    show "0 < ?d"
+      using d(1) assms(3) by auto
+    fix t :: real
+    assume as: "c \<le> t" "t < c + ?d"
+    have *: "integral {a .. c} f = integral {a .. b} f - integral {c .. b} f"
+      "integral {a .. t} f = integral {a .. b} f - integral {t .. b} f"
+      apply (simp_all only: algebra_simps)
+      apply (rule_tac[!] integral_combine)
+      using assms as
+      apply auto
+      done
+    have "(- c) - d < (- t) \<and> - t \<le> - c"
+      using as by auto note d(2)[rule_format,OF this]
+    then show "norm (integral {a .. c} f - integral {a .. t} f) < e"
+      unfolding *
+      unfolding integral_reflect
+      apply (subst norm_minus_commute)
+      apply (auto simp add: algebra_simps)
+      done
+  qed
+qed
+
+lemma indefinite_integral_continuous:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "f integrable_on {a .. b}"
+  shows "continuous_on {a .. b} (\<lambda>x. integral {a .. x} f)"
+proof (unfold continuous_on_iff, safe)
+  fix x e :: real
+  assume as: "x \<in> {a .. b}" "e > 0"
+  let ?thesis = "\<exists>d>0. \<forall>x'\<in>{a .. b}. dist x' x < d \<longrightarrow> dist (integral {a .. x'} f) (integral {a .. x} f) < e"
+  {
+    presume *: "a < b \<Longrightarrow> ?thesis"
+    show ?thesis
+      apply cases
+      apply (rule *)
+      apply assumption
+    proof goal_cases
+      case 1
+      then have "cbox a b = {x}"
+        using as(1)
+        apply -
+        apply (rule set_eqI)
+        apply auto
+        done
+      then show ?case using \<open>e > 0\<close> by auto
+    qed
+  }
+  assume "a < b"
+  have "(x = a \<or> x = b) \<or> (a < x \<and> x < b)"
+    using as(1) by auto
+  then show ?thesis
+    apply (elim disjE)
+  proof -
+    assume "x = a"
+    have "a \<le> a" ..
+    from indefinite_integral_continuous_right[OF assms(1) this \<open>a<b\<close> \<open>e>0\<close>] guess d . note d=this
+    show ?thesis
+      apply rule
+      apply rule
+      apply (rule d)
+      apply safe
+      apply (subst dist_commute)
+      unfolding \<open>x = a\<close> dist_norm
+      apply (rule d(2)[rule_format])
+      apply auto
+      done
+  next
+    assume "x = b"
+    have "b \<le> b" ..
+    from indefinite_integral_continuous_left[OF assms(1) \<open>a<b\<close> this \<open>e>0\<close>] guess d . note d=this
+    show ?thesis
+      apply rule
+      apply rule
+      apply (rule d)
+      apply safe
+      apply (subst dist_commute)
+      unfolding \<open>x = b\<close> dist_norm
+      apply (rule d(2)[rule_format])
+      apply auto
+      done
+  next
+    assume "a < x \<and> x < b"
+    then have xl: "a < x" "x \<le> b" and xr: "a \<le> x" "x < b"
+      by auto
+    from indefinite_integral_continuous_left [OF assms(1) xl \<open>e>0\<close>] guess d1 . note d1=this
+    from indefinite_integral_continuous_right[OF assms(1) xr \<open>e>0\<close>] guess d2 . note d2=this
+    show ?thesis
+      apply (rule_tac x="min d1 d2" in exI)
+    proof safe
+      show "0 < min d1 d2"
+        using d1 d2 by auto
+      fix y
+      assume "y \<in> {a .. b}" and "dist y x < min d1 d2"
+      then show "dist (integral {a .. y} f) (integral {a .. x} f) < e"
+        apply (subst dist_commute)
+        apply (cases "y < x")
+        unfolding dist_norm
+        apply (rule d1(2)[rule_format])
+        defer
+        apply (rule d2(2)[rule_format])
+        unfolding not_less
+        apply (auto simp add: field_simps)
+        done
+    qed
+  qed
+qed
+
+
+subsection \<open>This doesn't directly involve integration, but that gives an easy proof.\<close>
+
+lemma has_derivative_zero_unique_strong_interval:
+  fixes f :: "real \<Rightarrow> 'a::banach"
+  assumes "finite k"
+    and "continuous_on {a .. b} f"
+    and "f a = y"
+    and "\<forall>x\<in>({a .. b} - k). (f has_derivative (\<lambda>h. 0)) (at x within {a .. b})" "x \<in> {a .. b}"
+  shows "f x = y"
+proof -
+  have ab: "a \<le> b"
+    using assms by auto
+  have *: "a \<le> x"
+    using assms(5) by auto
+  have "((\<lambda>x. 0::'a) has_integral f x - f a) {a .. x}"
+    apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *])
+    apply (rule continuous_on_subset[OF assms(2)])
+    defer
+    apply safe
+    unfolding has_vector_derivative_def
+    apply (subst has_derivative_within_open[symmetric])
+    apply assumption
+    apply (rule open_greaterThanLessThan)
+    apply (rule has_derivative_within_subset[where s="{a .. b}"])
+    using assms(4) assms(5)
+    apply (auto simp: mem_box)
+    done
+  note this[unfolded *]
+  note has_integral_unique[OF has_integral_0 this]
+  then show ?thesis
+    unfolding assms by auto
+qed
+
+
+subsection \<open>Generalize a bit to any convex set.\<close>
+
+lemma has_derivative_zero_unique_strong_convex:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "convex s"
+    and "finite k"
+    and "continuous_on s f"
+    and "c \<in> s"
+    and "f c = y"
+    and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    and "x \<in> s"
+  shows "f x = y"
+proof -
+  {
+    presume *: "x \<noteq> c \<Longrightarrow> ?thesis"
+    show ?thesis
+      apply cases
+      apply (rule *)
+      apply assumption
+      unfolding assms(5)[symmetric]
+      apply auto
+      done
+  }
+  assume "x \<noteq> c"
+  note conv = assms(1)[unfolded convex_alt,rule_format]
+  have as1: "continuous_on {0 ..1} (f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x))"
+    apply (rule continuous_intros)+
+    apply (rule continuous_on_subset[OF assms(3)])
+    apply safe
+    apply (rule conv)
+    using assms(4,7)
+    apply auto
+    done
+  have *: "t = xa" if "(1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x" for t xa
+  proof -
+    from that have "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c"
+      unfolding scaleR_simps by (auto simp add: algebra_simps)
+    then show ?thesis
+      using \<open>x \<noteq> c\<close> by auto
+  qed
+  have as2: "finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}"
+    using assms(2)
+    apply (rule finite_surj[where f="\<lambda>z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"])
+    apply safe
+    unfolding image_iff
+    apply rule
+    defer
+    apply assumption
+    apply (rule sym)
+    apply (rule some_equality)
+    defer
+    apply (drule *)
+    apply auto
+    done
+  have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x)) 1 = y"
+    apply (rule has_derivative_zero_unique_strong_interval[OF as2 as1, of ])
+    unfolding o_def
+    using assms(5)
+    defer
+    apply -
+    apply rule
+  proof -
+    fix t
+    assume as: "t \<in> {0 .. 1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \<in> k}"
+    have *: "c - t *\<^sub>R c + t *\<^sub>R x \<in> s - k"
+      apply safe
+      apply (rule conv[unfolded scaleR_simps])
+      using \<open>x \<in> s\<close> \<open>c \<in> s\<close> as
+      by (auto simp add: algebra_simps)
+    have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x))
+      (at t within {0 .. 1})"
+      apply (intro derivative_eq_intros)
+      apply simp_all
+      apply (simp add: field_simps)
+      unfolding scaleR_simps
+      apply (rule has_derivative_within_subset,rule assms(6)[rule_format])
+      apply (rule *)
+      apply safe
+      apply (rule conv[unfolded scaleR_simps])
+      using \<open>x \<in> s\<close> \<open>c \<in> s\<close>
+      apply auto
+      done
+    then show "((\<lambda>xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\<lambda>h. 0)) (at t within {0 .. 1})"
+      unfolding o_def .
+  qed auto
+  then show ?thesis
+    by auto
+qed
+
+
+text \<open>Also to any open connected set with finite set of exceptions. Could
+ generalize to locally convex set with limpt-free set of exceptions.\<close>
+
+lemma has_derivative_zero_unique_strong_connected:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "connected s"
+    and "open s"
+    and "finite k"
+    and "continuous_on s f"
+    and "c \<in> s"
+    and "f c = y"
+    and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    and "x\<in>s"
+  shows "f x = y"
+proof -
+  have "{x \<in> s. f x \<in> {y}} = {} \<or> {x \<in> s. f x \<in> {y}} = s"
+    apply (rule assms(1)[unfolded connected_clopen,rule_format])
+    apply rule
+    defer
+    apply (rule continuous_closedin_preimage[OF assms(4) closed_singleton])
+    apply (rule open_openin_trans[OF assms(2)])
+    unfolding open_contains_ball
+  proof safe
+    fix x
+    assume "x \<in> s"
+    from assms(2)[unfolded open_contains_ball,rule_format,OF this] guess e .. note e=conjunctD2[OF this]
+    show "\<exists>e>0. ball x e \<subseteq> {xa \<in> s. f xa \<in> {f x}}"
+      apply rule
+      apply rule
+      apply (rule e)
+    proof safe
+      fix y
+      assume y: "y \<in> ball x e"
+      then show "y \<in> s"
+        using e by auto
+      show "f y = f x"
+        apply (rule has_derivative_zero_unique_strong_convex[OF convex_ball])
+        apply (rule assms)
+        apply (rule continuous_on_subset)
+        apply (rule assms)
+        apply (rule e)+
+        apply (subst centre_in_ball)
+        apply (rule e)
+        apply rule
+        apply safe
+        apply (rule has_derivative_within_subset)
+        apply (rule assms(7)[rule_format])
+        using y e
+        apply auto
+        done
+    qed
+  qed
+  then show ?thesis
+    using \<open>x \<in> s\<close> \<open>f c = y\<close> \<open>c \<in> s\<close> by auto
+qed
+
+lemma has_derivative_zero_connected_constant:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "connected s"
+      and "open s"
+      and "finite k"
+      and "continuous_on s f"
+      and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
+    obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c"
+proof (cases "s = {}")
+  case True
+  then show ?thesis
+by (metis empty_iff that)
+next
+  case False
+  then obtain c where "c \<in> s"
+    by (metis equals0I)
+  then show ?thesis
+    by (metis has_derivative_zero_unique_strong_connected assms that)
+qed
+
+
+subsection \<open>Integrating characteristic function of an interval\<close>
+
+lemma has_integral_restrict_open_subinterval:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "(f has_integral i) (cbox c d)"
+    and "cbox c d \<subseteq> cbox a b"
+  shows "((\<lambda>x. if x \<in> box c d then f x else 0) has_integral i) (cbox a b)"
+proof -
+  define g where [abs_def]: "g x = (if x \<in>box c d then f x else 0)" for x
+  {
+    presume *: "cbox c d \<noteq> {} \<Longrightarrow> ?thesis"
+    show ?thesis
+      apply cases
+      apply (rule *)
+      apply assumption
+    proof goal_cases
+      case prems: 1
+      then have *: "box c d = {}"
+        by (metis bot.extremum_uniqueI box_subset_cbox)
+      show ?thesis
+        using assms(1)
+        unfolding *
+        using prems
+        by auto
+    qed
+  }
+  assume "cbox c d \<noteq> {}"
+  from partial_division_extend_1[OF assms(2) this] guess p . note p=this
+  note operat = comm_monoid_set.operative_division[OF add.comm_monoid_set_lift_option operative_integral p(1), symmetric]
+  let ?P = "(if g integrable_on cbox a b then Some (integral (cbox a b) g) else None) = Some i"
+  {
+    presume "?P"
+    then have "g integrable_on cbox a b \<and> integral (cbox a b) g = i"
+      apply -
+      apply cases
+      apply (subst(asm) if_P)
+      apply assumption
+      apply auto
+      done
+    then show ?thesis
+      using integrable_integral
+      unfolding g_def
+      by auto
+  }
+  let ?F = "comm_monoid_set.F (lift_option op +) (Some 0)"
+  have iterate:"?F (\<lambda>i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0"
+  proof (intro comm_monoid_set.neutral[OF add.comm_monoid_set_lift_option] ballI)
+    fix x
+    assume x: "x \<in> p - {cbox c d}"
+    then have "x \<in> p"
+      by auto
+    note div = division_ofD(2-5)[OF p(1) this]
+    from div(3) guess u v by (elim exE) note uv=this
+    have "interior x \<inter> interior (cbox c d) = {}"
+      using div(4)[OF p(2)] x by auto
+    then have "(g has_integral 0) x"
+      unfolding uv
+      apply -
+      apply (rule has_integral_spike_interior[where f="\<lambda>x. 0"])
+      unfolding g_def interior_cbox
+      apply auto
+      done
+    then show "(if g integrable_on x then Some (integral x g) else None) = Some 0"
+      by auto
+  qed
+
+  have *: "p = insert (cbox c d) (p - {cbox c d})"
+    using p by auto
+  have **: "g integrable_on cbox c d"
+    apply (rule integrable_spike_interior[where f=f])
+    unfolding g_def  using assms(1)
+    apply auto
+    done
+  moreover
+  have "integral (cbox c d) g = i"
+    apply (rule has_integral_unique[OF _ assms(1)])
+    apply (rule has_integral_spike_interior[where f=g])
+    defer
+    apply (rule integrable_integral[OF **])
+    unfolding g_def
+    apply auto
+    done
+  ultimately show ?P
+    unfolding operat
+    using p
+    apply (subst *)
+    apply (subst comm_monoid_set.insert[OF add.comm_monoid_set_lift_option])
+    apply (simp_all add: division_of_finite iterate)
+    done
+qed
+
+lemma has_integral_restrict_closed_subinterval:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "(f has_integral i) (cbox c d)"
+    and "cbox c d \<subseteq> cbox a b"
+  shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)"
+proof -
+  note has_integral_restrict_open_subinterval[OF assms]
+  note * = has_integral_spike[OF negligible_frontier_interval _ this]
+  show ?thesis
+    apply (rule *[of c d])
+    using box_subset_cbox[of c d]
+    apply auto
+    done
+qed
+
+lemma has_integral_restrict_closed_subintervals_eq:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
+  assumes "cbox c d \<subseteq> cbox a b"
+  shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b) \<longleftrightarrow> (f has_integral i) (cbox c d)"
+  (is "?l = ?r")
+proof (cases "cbox c d = {}")
+  case False
+  let ?g = "\<lambda>x. if x \<in> cbox c d then f x else 0"
+  show ?thesis
+    apply rule
+    defer
+    apply (rule has_integral_restrict_closed_subinterval[OF _ assms])
+    apply assumption
+  proof -
+    assume ?l
+    then have "?g integrable_on cbox c d"
+      using assms has_integral_integrable integrable_subinterval by blast
+    then have *: "f integrable_on cbox c d"
+      apply -
+      apply (rule integrable_eq)
+      apply auto
+      done
+    then have "i = integral (cbox c d) f"
+      apply -
+      apply (rule has_integral_unique)
+      apply (rule \<open>?l\<close>)
+      apply (rule has_integral_restrict_closed_subinterval[OF _ assms])
+      apply auto
+      done
+    then show ?r
+      using * by auto
+  qed
+qed auto
+
+
+text \<open>Hence we can apply the limit process uniformly to all integrals.\<close>
+
+lemma has_integral':
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "(f has_integral i) s \<longleftrightarrow>
+    (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      (\<exists>z. ((\<lambda>x. if x \<in> s then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - i) < e))"
+  (is "?l \<longleftrightarrow> (\<forall>e>0. ?r e)")
+proof -
+  {
+    presume *: "\<exists>a b. s = cbox a b \<Longrightarrow> ?thesis"
+    show ?thesis
+      apply cases
+      apply (rule *)
+      apply assumption
+      apply (subst has_integral_alt)
+      apply auto
+      done
+  }
+  assume "\<exists>a b. s = cbox a b"
+  then guess a b by (elim exE) note s=this
+  from bounded_cbox[of a b, unfolded bounded_pos] guess B ..
+  note B = conjunctD2[OF this,rule_format] show ?thesis
+    apply safe
+  proof -
+    fix e :: real
+    assume ?l and "e > 0"
+    show "?r e"
+      apply (rule_tac x="B+1" in exI)
+      apply safe
+      defer
+      apply (rule_tac x=i in exI)
+    proof
+      fix c d :: 'n
+      assume as: "ball 0 (B+1) \<subseteq> cbox c d"
+      then show "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) (cbox c d)"
+        unfolding s
+        apply -
+        apply (rule has_integral_restrict_closed_subinterval)
+        apply (rule \<open>?l\<close>[unfolded s])
+        apply safe
+        apply (drule B(2)[rule_format])
+        unfolding subset_eq
+        apply (erule_tac x=x in ballE)
+        apply (auto simp add: dist_norm)
+        done
+    qed (insert B \<open>e>0\<close>, auto)
+  next
+    assume as: "\<forall>e>0. ?r e"
+    from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
+    define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
+    define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
+    have c_d: "cbox a b \<subseteq> cbox c d"
+      apply safe
+      apply (drule B(2))
+      unfolding mem_box
+    proof
+      fix x i
+      show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm x \<le> B" and "i \<in> Basis"
+        using that and Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
+        unfolding c_def d_def
+        by (auto simp add: field_simps setsum_negf)
+    qed
+    have "ball 0 C \<subseteq> cbox c d"
+      apply (rule subsetI)
+      unfolding mem_box mem_ball dist_norm
+    proof
+      fix x i :: 'n
+      assume x: "norm (0 - x) < C" and i: "i \<in> Basis"
+      show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
+        using Basis_le_norm[OF i, of x] and x i
+        unfolding c_def d_def
+        by (auto simp: setsum_negf)
+    qed
+    from C(2)[OF this] have "\<exists>y. (f has_integral y) (cbox a b)"
+      unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric]
+      unfolding s
+      by auto
+    then guess y .. note y=this
+
+    have "y = i"
+    proof (rule ccontr)
+      assume "\<not> ?thesis"
+      then have "0 < norm (y - i)"
+        by auto
+      from as[rule_format,OF this] guess C ..  note C=conjunctD2[OF this,rule_format]
+      define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
+      define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
+      have c_d: "cbox a b \<subseteq> cbox c d"
+        apply safe
+        apply (drule B(2))
+        unfolding mem_box
+      proof
+        fix x i :: 'n
+        assume "norm x \<le> B" and "i \<in> Basis"
+        then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
+          using Basis_le_norm[of i x]
+          unfolding c_def d_def
+          by (auto simp add: field_simps setsum_negf)
+      qed
+      have "ball 0 C \<subseteq> cbox c d"
+        apply (rule subsetI)
+        unfolding mem_box mem_ball dist_norm
+      proof
+        fix x i :: 'n
+        assume "norm (0 - x) < C" and "i \<in> Basis"
+        then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
+          using Basis_le_norm[of i x]
+          unfolding c_def d_def
+          by (auto simp: setsum_negf)
+      qed
+      note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s]
+      note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]]
+      then have "z = y" and "norm (z - i) < norm (y - i)"
+        apply -
+        apply (rule has_integral_unique[OF _ y(1)])
+        apply assumption
+        apply assumption
+        done
+      then show False
+        by auto
+    qed
+    then show ?l
+      using y
+      unfolding s
+      by auto
+  qed
+qed
+
+lemma has_integral_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "(f has_integral i) s"
+    and "(g has_integral j) s"
+    and "\<forall>x\<in>s. f x \<le> g x"
+  shows "i \<le> j"
+  using has_integral_component_le[OF _ assms(1-2), of 1]
+  using assms(3)
+  by auto
+
+lemma integral_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "f integrable_on s"
+    and "g integrable_on s"
+    and "\<forall>x\<in>s. f x \<le> g x"
+  shows "integral s f \<le> integral s g"
+  by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)])
+
+lemma has_integral_nonneg:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "(f has_integral i) s"
+    and "\<forall>x\<in>s. 0 \<le> f x"
+  shows "0 \<le> i"
+  using has_integral_component_nonneg[of 1 f i s]
+  unfolding o_def
+  using assms
+  by auto
+
+lemma integral_nonneg:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "f integrable_on s"
+    and "\<forall>x\<in>s. 0 \<le> f x"
+  shows "0 \<le> integral s f"
+  by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)])
+
+
+text \<open>Hence a general restriction property.\<close>
+
+lemma has_integral_restrict[simp]:
+  assumes "s \<subseteq> t"
+  shows "((\<lambda>x. if x \<in> s then f x else (0::'a::banach)) has_integral i) t \<longleftrightarrow> (f has_integral i) s"
+proof -
+  have *: "\<And>x. (if x \<in> t then if x \<in> s then f x else 0 else 0) =  (if x\<in>s then f x else 0)"
+    using assms by auto
+  show ?thesis
+    apply (subst(2) has_integral')
+    apply (subst has_integral')
+    unfolding *
+    apply rule
+    done
+qed
+
+lemma has_integral_restrict_univ:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) UNIV \<longleftrightarrow> (f has_integral i) s"
+  by auto
+
+lemma has_integral_on_superset:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
+    and "s \<subseteq> t"
+    and "(f has_integral i) s"
+  shows "(f has_integral i) t"
+proof -
+  have "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. if x \<in> t then f x else 0)"
+    apply rule
+    using assms(1-2)
+    apply auto
+    done
+  then show ?thesis
+    using assms(3)
+    apply (subst has_integral_restrict_univ[symmetric])
+    apply (subst(asm) has_integral_restrict_univ[symmetric])
+    apply auto
+    done
+qed
+
+lemma integrable_on_superset:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
+    and "s \<subseteq> t"
+    and "f integrable_on s"
+  shows "f integrable_on t"
+  using assms
+  unfolding integrable_on_def
+  by (auto intro:has_integral_on_superset)
+
+lemma integral_restrict_univ[intro]:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "f integrable_on s \<Longrightarrow> integral UNIV (\<lambda>x. if x \<in> s then f x else 0) = integral s f"
+  apply (rule integral_unique)
+  unfolding has_integral_restrict_univ
+  apply auto
+  done
+
+lemma integrable_restrict_univ:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "(\<lambda>x. if x \<in> s then f x else 0) integrable_on UNIV \<longleftrightarrow> f integrable_on s"
+  unfolding integrable_on_def
+  by auto
+
+lemma negligible_on_intervals: "negligible s \<longleftrightarrow> (\<forall>a b. negligible(s \<inter> cbox a b))" (is "?l \<longleftrightarrow> ?r")
+proof
+  assume ?r
+  show ?l
+    unfolding negligible_def
+  proof safe
+    fix a b
+    show "(indicator s has_integral 0) (cbox a b)"
+      apply (rule has_integral_negligible[OF \<open>?r\<close>[rule_format,of a b]])
+      unfolding indicator_def
+      apply auto
+      done
+  qed
+qed auto
+
+lemma negligible_translation:
+  assumes "negligible S"
+    shows "negligible (op + c ` S)"
+proof -
+  have inj: "inj (op + c)"
+    by simp
+  show ?thesis
+  using assms
+  proof (clarsimp simp: negligible_def)
+    fix a b
+    assume "\<forall>x y. (indicator S has_integral 0) (cbox x y)"
+    then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))"
+      by (meson Diff_iff assms has_integral_negligible indicator_simps(2))
+    have eq: "indicator (op + c ` S) = (\<lambda>x. indicator S (x - c))"
+      by (force simp add: indicator_def)
+    show "(indicator (op + c ` S) has_integral 0) (cbox a b)"
+      using has_integral_affinity [OF *, of 1 "-c"]
+            cbox_translation [of "c" "-c+a" "-c+b"]
+      by (simp add: eq add.commute)
+  qed
+qed
+
+lemma negligible_translation_rev:
+  assumes "negligible (op + c ` S)"
+    shows "negligible S"
+by (metis negligible_translation [OF assms, of "-c"] translation_galois)
+
+lemma has_integral_spike_set_eq:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "negligible ((s - t) \<union> (t - s))"
+  shows "(f has_integral y) s \<longleftrightarrow> (f has_integral y) t"
+  unfolding has_integral_restrict_univ[symmetric,of f]
+  apply (rule has_integral_spike_eq[OF assms])
+  by (auto split: if_split_asm)
+
+lemma has_integral_spike_set[dest]:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "negligible ((s - t) \<union> (t - s))"
+    and "(f has_integral y) s"
+  shows "(f has_integral y) t"
+  using assms has_integral_spike_set_eq
+  by auto
+
+lemma integrable_spike_set[dest]:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "negligible ((s - t) \<union> (t - s))"
+    and "f integrable_on s"
+  shows "f integrable_on t"
+  using assms(2)
+  unfolding integrable_on_def
+  unfolding has_integral_spike_set_eq[OF assms(1)] .
+
+lemma integrable_spike_set_eq:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "negligible ((s - t) \<union> (t - s))"
+  shows "f integrable_on s \<longleftrightarrow> f integrable_on t"
+  apply rule
+  apply (rule_tac[!] integrable_spike_set)
+  using assms
+  apply auto
+  done
+
+(*lemma integral_spike_set:
+ "\<forall>f:real^M->real^N g s t.
+        negligible(s DIFF t \<union> t DIFF s)
+        \<longrightarrow> integral s f = integral t f"
+qed  REPEAT STRIP_TAC THEN REWRITE_TAC[integral] THEN
+  AP_TERM_TAC THEN ABS_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
+  ASM_MESON_TAC[]);;
+
+lemma has_integral_interior:
+ "\<forall>f:real^M->real^N y s.
+        negligible(frontier s)
+        \<longrightarrow> ((f has_integral y) (interior s) \<longleftrightarrow> (f has_integral y) s)"
+qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
+  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
+    NEGLIGIBLE_SUBSET)) THEN
+  REWRITE_TAC[frontier] THEN
+  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
+  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
+  SET_TAC[]);;
+
+lemma has_integral_closure:
+ "\<forall>f:real^M->real^N y s.
+        negligible(frontier s)
+        \<longrightarrow> ((f has_integral y) (closure s) \<longleftrightarrow> (f has_integral y) s)"
+qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
+  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
+    NEGLIGIBLE_SUBSET)) THEN
+  REWRITE_TAC[frontier] THEN
+  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
+  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
+  SET_TAC[]);;*)
+
+
+subsection \<open>More lemmas that are useful later\<close>
+
+lemma has_integral_subset_component_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes k: "k \<in> Basis"
+    and as: "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)\<bullet>k"
+  shows "i\<bullet>k \<le> j\<bullet>k"
+proof -
+  note has_integral_restrict_univ[symmetric, of f]
+  note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this]
+  show ?thesis
+    apply (rule *)
+    using as(1,4)
+    apply auto
+    done
+qed
+
+lemma has_integral_subset_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "s \<subseteq> t"
+    and "(f has_integral i) s"
+    and "(f has_integral j) t"
+    and "\<forall>x\<in>t. 0 \<le> f x"
+  shows "i \<le> j"
+  using has_integral_subset_component_le[OF _ assms(1), of 1 f i j]
+  using assms
+  by auto
+
+lemma integral_subset_component_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "k \<in> Basis"
+    and "s \<subseteq> t"
+    and "f integrable_on s"
+    and "f integrable_on t"
+    and "\<forall>x \<in> t. 0 \<le> f x \<bullet> k"
+  shows "(integral s f)\<bullet>k \<le> (integral t f)\<bullet>k"
+  apply (rule has_integral_subset_component_le)
+  using assms
+  apply auto
+  done
+
+lemma integral_subset_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "s \<subseteq> t"
+    and "f integrable_on s"
+    and "f integrable_on t"
+    and "\<forall>x \<in> t. 0 \<le> f x"
+  shows "integral s f \<le> integral t f"
+  apply (rule has_integral_subset_le)
+  using assms
+  apply auto
+  done
+
+lemma has_integral_alt':
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "(f has_integral i) s \<longleftrightarrow> (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
+    (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e)"
+  (is "?l = ?r")
+proof
+  assume ?r
+  show ?l
+    apply (subst has_integral')
+    apply safe
+  proof goal_cases
+    case (1 e)
+    from \<open>?r\<close>[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this]
+    show ?case
+      apply rule
+      apply rule
+      apply (rule B)
+      apply safe
+      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0)" in exI)
+      apply (drule B(2)[rule_format])
+      using integrable_integral[OF \<open>?r\<close>[THEN conjunct1,rule_format]]
+      apply auto
+      done
+  qed
+next
+  assume ?l note as = this[unfolded has_integral'[of f],rule_format]
+  let ?f = "\<lambda>x. if x \<in> s then f x else 0"
+  show ?r
+  proof safe
+    fix a b :: 'n
+    from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format]
+    let ?a = "\<Sum>i\<in>Basis. min (a\<bullet>i) (-B) *\<^sub>R i::'n"
+    let ?b = "\<Sum>i\<in>Basis. max (b\<bullet>i) B *\<^sub>R i::'n"
+    show "?f integrable_on cbox a b"
+    proof (rule integrable_subinterval[of _ ?a ?b])
+      have "ball 0 B \<subseteq> cbox ?a ?b"
+        apply (rule subsetI)
+        unfolding mem_ball mem_box dist_norm
+      proof (rule, goal_cases)
+        case (1 x i)
+        then show ?case using Basis_le_norm[of i x]
+          by (auto simp add:field_simps)
+      qed
+      from B(2)[OF this] guess z .. note conjunct1[OF this]
+      then show "?f integrable_on cbox ?a ?b"
+        unfolding integrable_on_def by auto
+      show "cbox a b \<subseteq> cbox ?a ?b"
+        apply safe
+        unfolding mem_box
+        apply rule
+        apply (erule_tac x=i in ballE)
+        apply auto
+        done
+    qed
+
+    fix e :: real
+    assume "e > 0"
+    from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format]
+    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
+      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
+      apply rule
+      apply rule
+      apply (rule B)
+      apply safe
+    proof goal_cases
+      case 1
+      from B(2)[OF this] guess z .. note z=conjunctD2[OF this]
+      from integral_unique[OF this(1)] show ?case
+        using z(2) by auto
+    qed
+  qed
+qed
+
+
+subsection \<open>Continuity of the integral (for a 1-dimensional interval).\<close>
+
+lemma integrable_alt:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "f integrable_on s \<longleftrightarrow>
+    (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
+    (\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
+    norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) -
+      integral (cbox c d)  (\<lambda>x. if x \<in> s then f x else 0)) < e)"
+  (is "?l = ?r")
+proof
+  assume ?l
+  then guess y unfolding integrable_on_def .. note this[unfolded has_integral_alt'[of f]]
+  note y=conjunctD2[OF this,rule_format]
+  show ?r
+    apply safe
+    apply (rule y)
+  proof goal_cases
+    case (1 e)
+    then have "e/2 > 0"
+      by auto
+    from y(2)[OF this] guess B .. note B=conjunctD2[OF this,rule_format]
+    show ?case
+      apply rule
+      apply rule
+      apply (rule B)
+      apply safe
+    proof goal_cases
+      case prems: (1 a b c d)
+      show ?case
+        apply (rule norm_triangle_half_l)
+        using B(2)[OF prems(1)] B(2)[OF prems(2)]
+        apply auto
+        done
+    qed
+  qed
+next
+  assume ?r
+  note as = conjunctD2[OF this,rule_format]
+  let ?cube = "\<lambda>n. cbox (\<Sum>i\<in>Basis. - real n *\<^sub>R i::'n) (\<Sum>i\<in>Basis. real n *\<^sub>R i)"
+  have "Cauchy (\<lambda>n. integral (?cube n) (\<lambda>x. if x \<in> s then f x else 0))"
+  proof (unfold Cauchy_def, safe, goal_cases)
+    case (1 e)
+    from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format]
+    from real_arch_simple[of B] guess N .. note N = this
+    {
+      fix n
+      assume n: "n \<ge> N"
+      have "ball 0 B \<subseteq> ?cube n"
+        apply (rule subsetI)
+        unfolding mem_ball mem_box dist_norm
+      proof (rule, goal_cases)
+        case (1 x i)
+        then show ?case
+          using Basis_le_norm[of i x] \<open>i\<in>Basis\<close>
+          using n N
+          by (auto simp add: field_simps setsum_negf)
+      qed
+    }
+    then show ?case
+      apply -
+      apply (rule_tac x=N in exI)
+      apply safe
+      unfolding dist_norm
+      apply (rule B(2))
+      apply auto
+      done
+  qed
+  from this[unfolded convergent_eq_cauchy[symmetric]] guess i ..
+  note i = this[THEN LIMSEQ_D]
+
+  show ?l unfolding integrable_on_def has_integral_alt'[of f]
+    apply (rule_tac x=i in exI)
+    apply safe
+    apply (rule as(1)[unfolded integrable_on_def])
+  proof goal_cases
+    case (1 e)
+    then have *: "e/2 > 0" by auto
+    from i[OF this] guess N .. note N =this[rule_format]
+    from as(2)[OF *] guess B .. note B=conjunctD2[OF this,rule_format]
+    let ?B = "max (real N) B"
+    show ?case
+      apply (rule_tac x="?B" in exI)
+    proof safe
+      show "0 < ?B"
+        using B(1) by auto
+      fix a b :: 'n
+      assume ab: "ball 0 ?B \<subseteq> cbox a b"
+      from real_arch_simple[of ?B] guess n .. note n=this
+      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
+        apply (rule norm_triangle_half_l)
+        apply (rule B(2))
+        defer
+        apply (subst norm_minus_commute)
+        apply (rule N[of n])
+      proof safe
+        show "N \<le> n"
+          using n by auto
+        fix x :: 'n
+        assume x: "x \<in> ball 0 B"
+        then have "x \<in> ball 0 ?B"
+          by auto
+        then show "x \<in> cbox a b"
+          using ab by blast
+        show "x \<in> ?cube n"
+          using x
+          unfolding mem_box mem_ball dist_norm
+          apply -
+        proof (rule, goal_cases)
+          case (1 i)
+          then show ?case
+            using Basis_le_norm[of i x] \<open>i \<in> Basis\<close>
+            using n
+            by (auto simp add: field_simps setsum_negf)
+        qed
+      qed
+    qed
+  qed
+qed
+
+lemma integrable_altD:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on s"
+  shows "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b"
+    and "\<And>e. e > 0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
+      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d)  (\<lambda>x. if x \<in> s then f x else 0)) < e"
+  using assms[unfolded integrable_alt[of f]] by auto
+
+lemma integrable_on_subcbox:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on s"
+    and "cbox a b \<subseteq> s"
+  shows "f integrable_on cbox a b"
+  apply (rule integrable_eq)
+  defer
+  apply (rule integrable_altD(1)[OF assms(1)])
+  using assms(2)
+  apply auto
+  done
+
+
+subsection \<open>A straddling criterion for integrability\<close>
+
+lemma integrable_straddle_interval:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "\<forall>e>0. \<exists>g  h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and>
+    norm (i - j) < e \<and> (\<forall>x\<in>cbox a b. (g x) \<le> f x \<and> f x \<le> h x)"
+  shows "f integrable_on cbox a b"
+proof (subst integrable_cauchy, safe, goal_cases)
+  case (1 e)
+  then have e: "e/3 > 0"
+    by auto
+  note assms[rule_format,OF this]
+  then guess g h i j by (elim exE conjE) note obt = this
+  from obt(1)[unfolded has_integral[of g], rule_format, OF e] guess d1 .. note d1=conjunctD2[OF this,rule_format]
+  from obt(2)[unfolded has_integral[of h], rule_format, OF e] guess d2 .. note d2=conjunctD2[OF this,rule_format]
+  show ?case
+    apply (rule_tac x="\<lambda>x. d1 x \<inter> d2 x" in exI)
+    apply (rule conjI gauge_inter d1 d2)+
+    unfolding fine_inter
+  proof (safe, goal_cases)
+    have **: "\<And>i j g1 g2 h1 h2 f1 f2. g1 - h2 \<le> f1 - f2 \<Longrightarrow> f1 - f2 \<le> h1 - g2 \<Longrightarrow>
+      \<bar>i - j\<bar> < e / 3 \<Longrightarrow> \<bar>g2 - i\<bar> < e / 3 \<Longrightarrow> \<bar>g1 - i\<bar> < e / 3 \<Longrightarrow>
+      \<bar>h2 - j\<bar> < e / 3 \<Longrightarrow> \<bar>h1 - j\<bar> < e / 3 \<Longrightarrow> \<bar>f1 - f2\<bar> < e"
+    using \<open>e > 0\<close> by arith
+    case prems: (1 p1 p2)
+    note tagged_division_ofD(2-4) note * = this[OF prems(1)] this[OF prems(4)]
+
+    have "(\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R g x) \<ge> 0"
+      and "0 \<le> (\<Sum>(x, k)\<in>p2. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)"
+      and "(\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R g x) \<ge> 0"
+      and "0 \<le> (\<Sum>(x, k)\<in>p1. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x)"
+      unfolding setsum_subtractf[symmetric]
+      apply -
+      apply (rule_tac[!] setsum_nonneg)
+      apply safe
+      unfolding real_scaleR_def right_diff_distrib[symmetric]
+      apply (rule_tac[!] mult_nonneg_nonneg)
+    proof -
+      fix a b
+      assume ab: "(a, b) \<in> p1"
+      show "0 \<le> content b"
+        using *(3)[OF ab]
+        apply safe
+        apply (rule content_pos_le)
+        done
+      then show "0 \<le> content b" .
+      show "0 \<le> f a - g a" "0 \<le> h a - f a"
+        using *(1-2)[OF ab]
+        using obt(4)[rule_format,of a]
+        by auto
+    next
+      fix a b
+      assume ab: "(a, b) \<in> p2"
+      show "0 \<le> content b"
+        using *(6)[OF ab]
+        apply safe
+        apply (rule content_pos_le)
+        done
+      then show "0 \<le> content b" .
+      show "0 \<le> f a - g a" and "0 \<le> h a - f a"
+        using *(4-5)[OF ab] using obt(4)[rule_format,of a] by auto
+    qed
+    then show ?case
+      apply -
+      unfolding real_norm_def
+      apply (rule **)
+      defer
+      defer
+      unfolding real_norm_def[symmetric]
+      apply (rule obt(3))
+      apply (rule d1(2)[OF conjI[OF prems(4,5)]])
+      apply (rule d1(2)[OF conjI[OF prems(1,2)]])
+      apply (rule d2(2)[OF conjI[OF prems(4,6)]])
+      apply (rule d2(2)[OF conjI[OF prems(1,3)]])
+      apply auto
+      done
+  qed
+qed
+
+lemma integrable_straddle:
+  fixes f :: "'n::euclidean_space \<Rightarrow> real"
+  assumes "\<forall>e>0. \<exists>g h i j. (g has_integral i) s \<and> (h has_integral j) s \<and>
+    norm (i - j) < e \<and> (\<forall>x\<in>s. g x \<le> f x \<and> f x \<le> h x)"
+  shows "f integrable_on s"
+proof -
+  have "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b"
+  proof (rule integrable_straddle_interval, safe, goal_cases)
+    case (1 a b e)
+    then have *: "e/4 > 0"
+      by auto
+    from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this
+    note obt(1)[unfolded has_integral_alt'[of g]]
+    note conjunctD2[OF this, rule_format]
+    note g = this(1) and this(2)[OF *]
+    from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
+    note obt(2)[unfolded has_integral_alt'[of h]]
+    note conjunctD2[OF this, rule_format]
+    note h = this(1) and this(2)[OF *]
+    from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
+    define c :: 'n where "c = (\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max B1 B2)) *\<^sub>R i)"
+    define d :: 'n where "d = (\<Sum>i\<in>Basis. max (b\<bullet>i) (max B1 B2) *\<^sub>R i)"
+    have *: "ball 0 B1 \<subseteq> cbox c d" "ball 0 B2 \<subseteq> cbox c d"
+      apply safe
+      unfolding mem_ball mem_box dist_norm
+      apply (rule_tac[!] ballI)
+    proof goal_cases
+      case (1 x i)
+      then show ?case using Basis_le_norm[of i x]
+        unfolding c_def d_def by auto
+    next
+      case (2 x i)
+      then show ?case using Basis_le_norm[of i x]
+        unfolding c_def d_def by auto
+    qed
+    have **: "\<And>ch cg ag ah::real. norm (ah - ag) \<le> norm (ch - cg) \<Longrightarrow> norm (cg - i) < e / 4 \<Longrightarrow>
+      norm (ch - j) < e / 4 \<Longrightarrow> norm (ag - ah) < e"
+      using obt(3)
+      unfolding real_norm_def
+      by arith
+    show ?case
+      apply (rule_tac x="\<lambda>x. if x \<in> s then g x else 0" in exI)
+      apply (rule_tac x="\<lambda>x. if x \<in> s then h x else 0" in exI)
+      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)" in exI)
+      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then h x else 0)" in exI)
+      apply safe
+      apply (rule_tac[1-2] integrable_integral,rule g)
+      apply (rule h)
+      apply (rule **[OF _ B1(2)[OF *(1)] B2(2)[OF *(2)]])
+    proof -
+      have *: "\<And>x f g. (if x \<in> s then f x else 0) - (if x \<in> s then g x else 0) =
+        (if x \<in> s then f x - g x else (0::real))"
+        by auto
+      note ** = abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF h g]]
+      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then h x else 0) -
+          integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)) \<le>
+        norm (integral (cbox c d) (\<lambda>x. if x \<in> s then h x else 0) -
+          integral (cbox c d) (\<lambda>x. if x \<in> s then g x else 0))"
+        unfolding integral_diff[OF h g,symmetric] real_norm_def
+        apply (subst **)
+        defer
+        apply (subst **)
+        defer
+        apply (rule has_integral_subset_le)
+        defer
+        apply (rule integrable_integral integrable_diff h g)+
+      proof safe
+        fix x
+        assume "x \<in> cbox a b"
+        then show "x \<in> cbox c d"
+          unfolding mem_box c_def d_def
+          apply -
+          apply rule
+          apply (erule_tac x=i in ballE)
+          apply auto
+          done
+      qed (insert obt(4), auto)
+    qed (insert obt(4), auto)
+  qed
+  note interv = this
+
+  show ?thesis
+    unfolding integrable_alt[of f]
+    apply safe
+    apply (rule interv)
+  proof goal_cases
+    case (1 e)
+    then have *: "e/3 > 0"
+      by auto
+    from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this
+    note obt(1)[unfolded has_integral_alt'[of g]]
+    note conjunctD2[OF this, rule_format]
+    note g = this(1) and this(2)[OF *]
+    from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
+    note obt(2)[unfolded has_integral_alt'[of h]]
+    note conjunctD2[OF this, rule_format]
+    note h = this(1) and this(2)[OF *]
+    from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
+    show ?case
+      apply (rule_tac x="max B1 B2" in exI)
+      apply safe
+      apply (rule max.strict_coboundedI1)
+      apply (rule B1)
+    proof -
+      fix a b c d :: 'n
+      assume as: "ball 0 (max B1 B2) \<subseteq> cbox a b" "ball 0 (max B1 B2) \<subseteq> cbox c d"
+      have **: "ball 0 B1 \<subseteq> ball (0::'n) (max B1 B2)" "ball 0 B2 \<subseteq> ball (0::'n) (max B1 B2)"
+        by auto
+      have *: "\<And>ga gc ha hc fa fc::real.
+        \<bar>ga - i\<bar> < e / 3 \<and> \<bar>gc - i\<bar> < e / 3 \<and> \<bar>ha - j\<bar> < e / 3 \<and>
+        \<bar>hc - j\<bar> < e / 3 \<and> \<bar>i - j\<bar> < e / 3 \<and> ga \<le> fa \<and> fa \<le> ha \<and> gc \<le> fc \<and> fc \<le> hc \<Longrightarrow>
+        \<bar>fa - fc\<bar> < e"
+        by (simp add: abs_real_def split: if_split_asm)
+      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d)
+        (\<lambda>x. if x \<in> s then f x else 0)) < e"
+        unfolding real_norm_def
+        apply (rule *)
+        apply safe
+        unfolding real_norm_def[symmetric]
+        apply (rule B1(2))
+        apply (rule order_trans)
+        apply (rule **)
+        apply (rule as(1))
+        apply (rule B1(2))
+        apply (rule order_trans)
+        apply (rule **)
+        apply (rule as(2))
+        apply (rule B2(2))
+        apply (rule order_trans)
+        apply (rule **)
+        apply (rule as(1))
+        apply (rule B2(2))
+        apply (rule order_trans)
+        apply (rule **)
+        apply (rule as(2))
+        apply (rule obt)
+        apply (rule_tac[!] integral_le)
+        using obt
+        apply (auto intro!: h g interv)
+        done
+    qed
+  qed
+qed
+
+
+subsection \<open>Adding integrals over several sets\<close>
+
+lemma has_integral_union:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "(f has_integral i) s"
+    and "(f has_integral j) t"
+    and "negligible (s \<inter> t)"
+  shows "(f has_integral (i + j)) (s \<union> t)"
+proof -
+  note * = has_integral_restrict_univ[symmetric, of f]
+  show ?thesis
+    unfolding *
+    apply (rule has_integral_spike[OF assms(3)])
+    defer
+    apply (rule has_integral_add[OF assms(1-2)[unfolded *]])
+    apply auto
+    done
+qed
+
+lemma integrable_union:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
+  assumes "negligible (A \<inter> B)" "f integrable_on A" "f integrable_on B"
+  shows   "f integrable_on (A \<union> B)"
+proof -
+  from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B"
+     by (auto simp: integrable_on_def)
+  from has_integral_union[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def)
+qed
+
+lemma integrable_union':
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
+  assumes "f integrable_on A" "f integrable_on B" "negligible (A \<inter> B)" "C = A \<union> B"
+  shows   "f integrable_on C"
+  using integrable_union[of A B f] assms by simp
+
+lemma has_integral_unions:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "finite t"
+    and "\<forall>s\<in>t. (f has_integral (i s)) s"
+    and "\<forall>s\<in>t. \<forall>s'\<in>t. s \<noteq> s' \<longrightarrow> negligible (s \<inter> s')"
+  shows "(f has_integral (setsum i t)) (\<Union>t)"
+proof -
+  note * = has_integral_restrict_univ[symmetric, of f]
+  have **: "negligible (\<Union>((\<lambda>(a,b). a \<inter> b) ` {(a,b). a \<in> t \<and> b \<in> {y. y \<in> t \<and> a \<noteq> y}}))"
+    apply (rule negligible_Union)
+    apply (rule finite_imageI)
+    apply (rule finite_subset[of _ "t \<times> t"])
+    defer
+    apply (rule finite_cartesian_product[OF assms(1,1)])
+    using assms(3)
+    apply auto
+    done
+  note assms(2)[unfolded *]
+  note has_integral_setsum[OF assms(1) this]
+  then show ?thesis
+    unfolding *
+    apply -
+    apply (rule has_integral_spike[OF **])
+    defer
+    apply assumption
+    apply safe
+  proof goal_cases
+    case prems: (1 x)
+    then show ?case
+    proof (cases "x \<in> \<Union>t")
+      case True
+      then guess s unfolding Union_iff .. note s=this
+      then have *: "\<forall>b\<in>t. x \<in> b \<longleftrightarrow> b = s"
+        using prems(3) by blast
+      show ?thesis
+        unfolding if_P[OF True]
+        apply (rule trans)
+        defer
+        apply (rule setsum.cong)
+        apply (rule refl)
+        apply (subst *)
+        apply assumption
+        apply (rule refl)
+        unfolding setsum.delta[OF assms(1)]
+        using s
+        apply auto
+        done
+    qed auto
+  qed
+qed
+
+
+text \<open>In particular adding integrals over a division, maybe not of an interval.\<close>
+
+lemma has_integral_combine_division:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "d division_of s"
+    and "\<forall>k\<in>d. (f has_integral (i k)) k"
+  shows "(f has_integral (setsum i d)) s"
+proof -
+  note d = division_ofD[OF assms(1)]
+  show ?thesis
+    unfolding d(6)[symmetric]
+    apply (rule has_integral_unions)
+    apply (rule d assms)+
+    apply rule
+    apply rule
+    apply rule
+  proof goal_cases
+    case prems: (1 s s')
+    from d(4)[OF this(1)] d(4)[OF this(2)] guess a c b d by (elim exE) note obt=this
+    from d(5)[OF prems] show ?case
+      unfolding obt interior_cbox
+      apply -
+      apply (rule negligible_subset[of "(cbox a b-box a b) \<union> (cbox c d-box c d)"])
+      apply (rule negligible_Un negligible_frontier_interval)+
+      apply auto
+      done
+  qed
+qed
+
+lemma integral_combine_division_bottomup:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "d division_of s"
+    and "\<forall>k\<in>d. f integrable_on k"
+  shows "integral s f = setsum (\<lambda>i. integral i f) d"
+  apply (rule integral_unique)
+  apply (rule has_integral_combine_division[OF assms(1)])
+  using assms(2)
+  unfolding has_integral_integral
+  apply assumption
+  done
+
+lemma has_integral_combine_division_topdown:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on s"
+    and "d division_of k"
+    and "k \<subseteq> s"
+  shows "(f has_integral (setsum (\<lambda>i. integral i f) d)) k"
+  apply (rule has_integral_combine_division[OF assms(2)])
+  apply safe
+  unfolding has_integral_integral[symmetric]
+proof goal_cases
+  case (1 k)
+  from division_ofD(2,4)[OF assms(2) this]
+  show ?case
+    apply safe
+    apply (rule integrable_on_subcbox)
+    apply (rule assms)
+    using assms(3)
+    apply auto
+    done
+qed
+
+lemma integral_combine_division_topdown:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on s"
+    and "d division_of s"
+  shows "integral s f = setsum (\<lambda>i. integral i f) d"
+  apply (rule integral_unique)
+  apply (rule has_integral_combine_division_topdown)
+  using assms
+  apply auto
+  done
+
+lemma integrable_combine_division:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "d division_of s"
+    and "\<forall>i\<in>d. f integrable_on i"
+  shows "f integrable_on s"
+  using assms(2)
+  unfolding integrable_on_def
+  by (metis has_integral_combine_division[OF assms(1)])
+
+lemma integrable_on_subdivision:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "d division_of i"
+    and "f integrable_on s"
+    and "i \<subseteq> s"
+  shows "f integrable_on i"
+  apply (rule integrable_combine_division assms)+
+  apply safe
+proof goal_cases
+  case 1
+  note division_ofD(2,4)[OF assms(1) this]
+  then show ?case
+    apply safe
+    apply (rule integrable_on_subcbox[OF assms(2)])
+    using assms(3)
+    apply auto
+    done
+qed
+
+
+subsection \<open>Also tagged divisions\<close>
+
+lemma has_integral_combine_tagged_division:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "p tagged_division_of s"
+    and "\<forall>(x,k) \<in> p. (f has_integral (i k)) k"
+  shows "(f has_integral (setsum (\<lambda>(x,k). i k) p)) s"
+proof -
+  have *: "(f has_integral (setsum (\<lambda>k. integral k f) (snd ` p))) s"
+    apply (rule has_integral_combine_division)
+    apply (rule division_of_tagged_division[OF assms(1)])
+    using assms(2)
+    unfolding has_integral_integral[symmetric]
+    apply safe
+    apply auto
+    done
+  then show ?thesis
+    apply -
+    apply (rule subst[where P="\<lambda>i. (f has_integral i) s"])
+    defer
+    apply assumption
+    apply (rule trans[of _ "setsum (\<lambda>(x,k). integral k f) p"])
+    apply (subst eq_commute)
+    apply (rule setsum.over_tagged_division_lemma[OF assms(1)])
+    apply (rule integral_null)
+    apply assumption
+    apply (rule setsum.cong)
+    using assms(2)
+    apply auto
+    done
+qed
+
+lemma integral_combine_tagged_division_bottomup:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "p tagged_division_of (cbox a b)"
+    and "\<forall>(x,k)\<in>p. f integrable_on k"
+  shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
+  apply (rule integral_unique)
+  apply (rule has_integral_combine_tagged_division[OF assms(1)])
+  using assms(2)
+  apply auto
+  done
+
+lemma has_integral_combine_tagged_division_topdown:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on cbox a b"
+    and "p tagged_division_of (cbox a b)"
+  shows "(f has_integral (setsum (\<lambda>(x,k). integral k f) p)) (cbox a b)"
+  apply (rule has_integral_combine_tagged_division[OF assms(2)])
+  apply safe
+proof goal_cases
+  case 1
+  note tagged_division_ofD(3-4)[OF assms(2) this]
+  then show ?case
+    using integrable_subinterval[OF assms(1)] by blast
+qed
+
+lemma integral_combine_tagged_division_topdown:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on cbox a b"
+    and "p tagged_division_of (cbox a b)"
+  shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
+  apply (rule integral_unique)
+  apply (rule has_integral_combine_tagged_division_topdown)
+  using assms
+  apply auto
+  done
+
+
+subsection \<open>Henstock's lemma\<close>
+
+lemma henstock_lemma_part1:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on cbox a b"
+    and "e > 0"
+    and "gauge d"
+    and "(\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)"
+    and p: "p tagged_partial_division_of (cbox a b)" "d fine p"
+  shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x - integral k f) p) \<le> e"
+  (is "?x \<le> e")
+proof -
+  { presume "\<And>k. 0<k \<Longrightarrow> ?x \<le> e + k" then show ?thesis by (blast intro: field_le_epsilon) }
+  fix k :: real
+  assume k: "k > 0"
+  note p' = tagged_partial_division_ofD[OF p(1)]
+  have "\<Union>(snd ` p) \<subseteq> cbox a b"
+    using p'(3) by fastforce
+  note partial_division_of_tagged_division[OF p(1)] this
+  from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)]
+  define r where "r = q - snd ` p"
+  have "snd ` p \<inter> r = {}"
+    unfolding r_def by auto
+  have r: "finite r"
+    using q' unfolding r_def by auto
+
+  have "\<forall>i\<in>r. \<exists>p. p tagged_division_of i \<and> d fine p \<and>
+    norm (setsum (\<lambda>(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)"
+    apply safe
+  proof goal_cases
+    case (1 i)
+    then have i: "i \<in> q"
+      unfolding r_def by auto
+    from q'(4)[OF this] guess u v by (elim exE) note uv=this
+    have *: "k / (real (card r) + 1) > 0" using k by simp
+    have "f integrable_on cbox u v"
+      apply (rule integrable_subinterval[OF assms(1)])
+      using q'(2)[OF i]
+      unfolding uv
+      apply auto
+      done
+    note integrable_integral[OF this, unfolded has_integral[of f]]
+    from this[rule_format,OF *] guess dd .. note dd=conjunctD2[OF this,rule_format]
+    note gauge_inter[OF \<open>gauge d\<close> dd(1)]
+    from fine_division_exists[OF this,of u v] guess qq .
+    then show ?case
+      apply (rule_tac x=qq in exI)
+      using dd(2)[of qq]
+      unfolding fine_inter uv
+      apply auto
+      done
+  qed
+  from bchoice[OF this] guess qq .. note qq=this[rule_format]
+
+  let ?p = "p \<union> \<Union>(qq ` r)"
+  have "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e"
+    apply (rule assms(4)[rule_format])
+  proof
+    show "d fine ?p"
+      apply (rule fine_union)
+      apply (rule p)
+      apply (rule fine_unions)
+      using qq
+      apply auto
+      done
+    note * = tagged_partial_division_of_union_self[OF p(1)]
+    have "p \<union> \<Union>(qq ` r) tagged_division_of \<Union>(snd ` p) \<union> \<Union>r"
+      using r
+    proof (rule tagged_division_union[OF * tagged_division_unions], goal_cases)
+      case 1
+      then show ?case
+        using qq by auto
+    next
+      case 2
+      then show ?case
+        apply rule
+        apply rule
+        apply rule
+        apply(rule q'(5))
+        unfolding r_def
+        apply auto
+        done
+    next
+      case 3
+      then show ?case
+        apply (rule inter_interior_unions_intervals)
+        apply fact
+        apply rule
+        apply rule
+        apply (rule q')
+        defer
+        apply rule
+        apply (subst Int_commute)
+        apply (rule inter_interior_unions_intervals)
+        apply (rule finite_imageI)
+        apply (rule p')
+        apply rule
+        defer
+        apply rule
+        apply (rule q')
+        using q(1) p'
+        unfolding r_def
+        apply auto
+        done
+    qed
+    moreover have "\<Union>(snd ` p) \<union> \<Union>r = cbox a b" and "{qq i |i. i \<in> r} = qq ` r"
+      unfolding Union_Un_distrib[symmetric] r_def
+      using q
+      by auto
+    ultimately show "?p tagged_division_of (cbox a b)"
+      by fastforce
+  qed
+
+  then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>\<Union>(qq ` r). content k *\<^sub>R f x) -
+    integral (cbox a b) f) < e"
+    apply (subst setsum.union_inter_neutral[symmetric])
+    apply (rule p')
+    prefer 3
+    apply assumption
+    apply rule
+    apply (rule r)
+    apply safe
+    apply (drule qq)
+  proof -
+    fix x l k
+    assume as: "(x, l) \<in> p" "(x, l) \<in> qq k" "k \<in> r"
+    note qq[OF this(3)]
+    note tagged_division_ofD(3,4)[OF conjunct1[OF this] as(2)]
+    from this(2) guess u v by (elim exE) note uv=this
+    have "l\<in>snd ` p" unfolding image_iff apply(rule_tac x="(x,l)" in bexI) using as by auto
+    then have "l \<in> q" "k \<in> q" "l \<noteq> k"
+      using as(1,3) q(1) unfolding r_def by auto
+    note q'(5)[OF this]
+    then have "interior l = {}"
+      using interior_mono[OF \<open>l \<subseteq> k\<close>] by blast
+    then show "content l *\<^sub>R f x = 0"
+      unfolding uv content_eq_0_interior[symmetric] by auto
+  qed auto
+
+  then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x))
+    (qq ` r) - integral (cbox a b) f) < e"
+    apply (subst (asm) setsum.Union_comp)
+    prefer 2
+    unfolding split_paired_all split_conv image_iff
+    apply (erule bexE)+
+  proof -
+    fix x m k l T1 T2
+    assume "(x, m) \<in> T1" "(x, m) \<in> T2" "T1 \<noteq> T2" "k \<in> r" "l \<in> r" "T1 = qq k" "T2 = qq l"
+    note as = this(1-5)[unfolded this(6-)]
+    note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]]
+    from this(2)[OF as(4,1)] guess u v by (elim exE) note uv=this
+    have *: "interior (k \<inter> l) = {}"
+      by (metis DiffE \<open>T1 = qq k\<close> \<open>T1 \<noteq> T2\<close> \<open>T2 = qq l\<close> as(4) as(5) interior_Int q'(5) r_def)
+    have "interior m = {}"
+      unfolding subset_empty[symmetric]
+      unfolding *[symmetric]
+      apply (rule interior_mono)
+      using kl(1)[OF as(4,1)] kl(1)[OF as(5,2)]
+      apply auto
+      done
+    then show "content m *\<^sub>R f x = 0"
+      unfolding uv content_eq_0_interior[symmetric]
+      by auto
+  qed (insert qq, auto)
+
+  then have **: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x) \<circ> qq) r -
+    integral (cbox a b) f) < e"
+    apply (subst (asm) setsum.reindex_nontrivial)
+    apply fact
+    apply (rule setsum.neutral)
+    apply rule
+    unfolding split_paired_all split_conv
+    defer
+    apply assumption
+  proof -
+    fix k l x m
+    assume as: "k \<in> r" "l \<in> r" "k \<noteq> l" "qq k = qq l" "(x, m) \<in> qq k"
+    note tagged_division_ofD(6)[OF qq[THEN conjunct1]]
+    from this[OF as(1)] this[OF as(2)] show "content m *\<^sub>R f x = 0"
+      using as(3) unfolding as by auto
+  qed
+
+  have *: "norm (cp - ip) \<le> e + k"
+    if "norm ((cp + cr) - i) < e"
+    and "norm (cr - ir) < k"
+    and "ip + ir = i"
+    for ir ip i cr cp
+  proof -
+    from that show ?thesis
+      using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]
+      unfolding that(3)[symmetric] norm_minus_cancel
+      by (auto simp add: algebra_simps)
+  qed
+
+  have "?x =  norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
+    unfolding split_def setsum_subtractf ..
+  also have "\<dots> \<le> e + k"
+    apply (rule *[OF **, where ir1="setsum (\<lambda>k. integral k f) r"])
+  proof goal_cases
+    case 1
+    have *: "k * real (card r) / (1 + real (card r)) < k"
+      using k by (auto simp add: field_simps)
+    show ?case
+      apply (rule le_less_trans[of _ "setsum (\<lambda>x. k / (real (card r) + 1)) r"])
+      unfolding setsum_subtractf[symmetric]
+      apply (rule setsum_norm_le)
+      apply rule
+      apply (drule qq)
+      defer
+      unfolding divide_inverse setsum_left_distrib[symmetric]
+      unfolding divide_inverse[symmetric]
+      using * apply (auto simp add: field_simps)
+      done
+  next
+    case 2
+    have *: "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)"
+      apply (subst setsum.reindex_nontrivial)
+      apply fact
+      unfolding split_paired_all snd_conv split_def o_def
+    proof -
+      fix x l y m
+      assume as: "(x, l) \<in> p" "(y, m) \<in> p" "(x, l) \<noteq> (y, m)" "l = m"
+      from p'(4)[OF as(1)] guess u v by (elim exE) note uv=this
+      show "integral l f = 0"
+        unfolding uv
+        apply (rule integral_unique)
+        apply (rule has_integral_null)
+        unfolding content_eq_0_interior
+        using p'(5)[OF as(1-3)]
+        unfolding uv as(4)[symmetric]
+        apply auto
+        done
+    qed auto
+    from q(1) have **: "snd ` p \<union> q = q" by auto
+    show ?case
+      unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def
+      using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\<lambda>k. integral k f", symmetric]
+        by simp
+  qed
+  finally show "?x \<le> e + k" .
+qed
+
+lemma henstock_lemma_part2:
+  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
+  assumes "f integrable_on cbox a b"
+    and "e > 0"
+    and "gauge d"
+    and "\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e"
+    and "p tagged_partial_division_of (cbox a b)"
+    and "d fine p"
+  shows "setsum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
+  unfolding split_def
+  apply (rule setsum_norm_allsubsets_bound)
+  defer
+  apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)])
+  apply safe
+  apply (rule assms[rule_format,unfolded split_def])
+  defer
+  apply (rule tagged_partial_division_subset)
+  apply (rule assms)
+  apply assumption
+  apply (rule fine_subset)
+  apply assumption
+  apply (rule assms)
+  using assms(5)
+  apply auto
+  done
+
+lemma henstock_lemma:
+  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
+  assumes "f integrable_on cbox a b"
+    and "e > 0"
+  obtains d where "gauge d"
+    and "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+      setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
+proof -
+  have *: "e / (2 * (real DIM('n) + 1)) > 0" using assms(2) by simp
+  from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this]
+  guess d .. note d = conjunctD2[OF this]
+  show thesis
+    apply (rule that)
+    apply (rule d)
+  proof (safe, goal_cases)
+    case (1 p)
+    note * = henstock_lemma_part2[OF assms(1) * d this]
+    show ?case
+      apply (rule le_less_trans[OF *])
+      using \<open>e > 0\<close>
+      apply (auto simp add: field_simps)
+      done
+  qed
+qed
+
+
+subsection \<open>Geometric progression\<close>
+
+text \<open>FIXME: Should one or more of these theorems be moved to @{file
+"~~/src/HOL/Set_Interval.thy"}, alongside \<open>geometric_sum\<close>?\<close>
+
+lemma sum_gp_basic:
+  fixes x :: "'a::ring_1"
+  shows "(1 - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
+proof -
+  define y where "y = 1 - x"
+  have "y * (\<Sum>i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n"
+    by (induct n) (simp_all add: algebra_simps)
+  then show ?thesis
+    unfolding y_def by simp
+qed
+
+lemma sum_gp_multiplied:
+  assumes mn: "m \<le> n"
+  shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
+  (is "?lhs = ?rhs")
+proof -
+  let ?S = "{0..(n - m)}"
+  from mn have mn': "n - m \<ge> 0"
+    by arith
+  let ?f = "op + m"
+  have i: "inj_on ?f ?S"
+    unfolding inj_on_def by auto
+  have f: "?f ` ?S = {m..n}"
+    using mn
+    apply (auto simp add: image_iff Bex_def)
+    apply presburger
+    done
+  have th: "op ^ x \<circ> op + m = (\<lambda>i. x^m * x^i)"
+    by (rule ext) (simp add: power_add power_mult)
+  from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
+  have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})"
+    by simp
+  then show ?thesis
+    unfolding sum_gp_basic
+    using mn
+    by (simp add: field_simps power_add[symmetric])
+qed
+
+lemma sum_gp:
+  "setsum (op ^ (x::'a::{field})) {m .. n} =
+    (if n < m then 0
+     else if x = 1 then of_nat ((n + 1) - m)
+     else (x^ m - x^ (Suc n)) / (1 - x))"
+proof -
+  {
+    assume nm: "n < m"
+    then have ?thesis by simp
+  }
+  moreover
+  {
+    assume "\<not> n < m"
+    then have nm: "m \<le> n"
+      by arith
+    {
+      assume x: "x = 1"
+      then have ?thesis
+        by simp
+    }
+    moreover
+    {
+      assume x: "x \<noteq> 1"
+      then have nz: "1 - x \<noteq> 0"
+        by simp
+      from sum_gp_multiplied[OF nm, of x] nz have ?thesis
+        by (simp add: field_simps)
+    }
+    ultimately have ?thesis by blast
+  }
+  ultimately show ?thesis by blast
+qed
+
+lemma sum_gp_offset:
+  "setsum (op ^ (x::'a::{field})) {m .. m+n} =
+    (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
+  unfolding sum_gp[of x m "m + n"] power_Suc
+  by (simp add: field_simps power_add)
+
+
+subsection \<open>Monotone convergence (bounded interval first)\<close>
+
+lemma monotone_convergence_interval:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
+  assumes "\<forall>k. (f k) integrable_on cbox a b"
+    and "\<forall>k. \<forall>x\<in>cbox a b.(f k x) \<le> f (Suc k) x"
+    and "\<forall>x\<in>cbox a b. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+    and "bounded {integral (cbox a b) (f k) | k . k \<in> UNIV}"
+  shows "g integrable_on cbox a b \<and> ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> integral (cbox a b) g) sequentially"
+proof (cases "content (cbox a b) = 0")
+  case True
+  show ?thesis
+    using integrable_on_null[OF True]
+    unfolding integral_null[OF True]
+    using tendsto_const
+    by auto
+next
+  case False
+  have fg: "\<forall>x\<in>cbox a b. \<forall>k. (f k x) \<bullet> 1 \<le> (g x) \<bullet> 1"
+  proof safe
+    fix x k
+    assume x: "x \<in> cbox a b"
+    note * = Lim_component_ge[OF assms(3)[rule_format, OF x] trivial_limit_sequentially]
+    show "f k x \<bullet> 1 \<le> g x \<bullet> 1"
+      apply (rule *)
+      unfolding eventually_sequentially
+      apply (rule_tac x=k in exI)
+      apply -
+      apply (rule transitive_stepwise_le)
+      using assms(2)[rule_format, OF x]
+      apply auto
+      done
+  qed
+  have "\<exists>i. ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> i) sequentially"
+    apply (rule bounded_increasing_convergent)
+    defer
+    apply rule
+    apply (rule integral_le)
+    apply safe
+    apply (rule assms(1-2)[rule_format])+
+    using assms(4)
+    apply auto
+    done
+  then guess i .. note i=this
+  have i': "\<And>k. (integral(cbox a b) (f k)) \<le> i\<bullet>1"
+    apply (rule Lim_component_ge)
+    apply (rule i)
+    apply (rule trivial_limit_sequentially)
+    unfolding eventually_sequentially
+    apply (rule_tac x=k in exI)
+    apply (rule transitive_stepwise_le)
+    prefer 3
+    unfolding inner_simps real_inner_1_right
+    apply (rule integral_le)
+    apply (rule assms(1-2)[rule_format])+
+    using assms(2)
+    apply auto
+    done
+
+  have "(g has_integral i) (cbox a b)"
+    unfolding has_integral
+  proof (safe, goal_cases)
+    case e: (1 e)
+    then have "\<forall>k. (\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
+      norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e / 2 ^ (k + 2)))"
+      apply -
+      apply rule
+      apply (rule assms(1)[unfolded has_integral_integral has_integral,rule_format])
+      apply auto
+      done
+    from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format]
+
+    have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i\<bullet>1 - (integral (cbox a b) (f k)) \<and> i\<bullet>1 - (integral (cbox a b) (f k)) < e / 4"
+    proof -
+      have "e/4 > 0"
+        using e by auto
+      from LIMSEQ_D [OF i this] guess r ..
+      then show ?thesis
+        apply (rule_tac x=r in exI)
+        apply rule
+        apply (erule_tac x=k in allE)
+        subgoal for k using i'[of k] by auto
+        done
+    qed
+    then guess r .. note r=conjunctD2[OF this[rule_format]]
+
+    have "\<forall>x\<in>cbox a b. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)\<bullet>1 - (f k x)\<bullet>1 \<and>
+      (g x)\<bullet>1 - (f k x)\<bullet>1 < e / (4 * content(cbox a b))"
+    proof (rule, goal_cases)
+      case prems: (1 x)
+      have "e / (4 * content (cbox a b)) > 0"
+        using \<open>e>0\<close> False content_pos_le[of a b] by auto
+      from assms(3)[rule_format, OF prems, THEN LIMSEQ_D, OF this]
+      guess n .. note n=this
+      then show ?case
+        apply (rule_tac x="n + r" in exI)
+        apply safe
+        apply (erule_tac[2-3] x=k in allE)
+        unfolding dist_real_def
+        using fg[rule_format, OF prems]
+        apply (auto simp add: field_simps)
+        done
+    qed
+    from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format]
+    define d where "d x = c (m x) x" for x
+    show ?case
+      apply (rule_tac x=d in exI)
+    proof safe
+      show "gauge d"
+        using c(1) unfolding gauge_def d_def by auto
+    next
+      fix p
+      assume p: "p tagged_division_of (cbox a b)" "d fine p"
+      note p'=tagged_division_ofD[OF p(1)]
+      have "\<exists>a. \<forall>x\<in>p. m (fst x) \<le> a"
+        by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI)
+      then guess s .. note s=this
+      have *: "\<forall>a b c d. norm(a - b) \<le> e / 4 \<and> norm(b - c) < e / 2 \<and>
+        norm (c - d) < e / 4 \<longrightarrow> norm (a - d) < e"
+      proof (safe, goal_cases)
+        case (1 a b c d)
+        then show ?case
+          using norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
+            norm_triangle_lt[of "a - b + (b - c)" "c - d" e]
+          unfolding norm_minus_cancel
+          by (auto simp add: algebra_simps)
+      qed
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - i) < e"
+        apply (rule *[rule_format,where
+          b="\<Sum>(x, k)\<in>p. content k *\<^sub>R f (m x) x" and c="\<Sum>(x, k)\<in>p. integral k (f (m x))"])
+      proof (safe, goal_cases)
+        case 1
+        show ?case
+          apply (rule order_trans[of _ "\<Sum>(x, k)\<in>p. content k * (e / (4 * content (cbox a b)))"])
+          unfolding setsum_subtractf[symmetric]
+          apply (rule order_trans)
+          apply (rule norm_setsum)
+          apply (rule setsum_mono)
+          unfolding split_paired_all split_conv
+          unfolding split_def setsum_left_distrib[symmetric] scaleR_diff_right[symmetric]
+          unfolding additive_content_tagged_division[OF p(1), unfolded split_def]
+        proof -
+          fix x k
+          assume xk: "(x, k) \<in> p"
+          then have x: "x \<in> cbox a b"
+            using p'(2-3)[OF xk] by auto
+          from p'(4)[OF xk] guess u v by (elim exE) note uv=this
+          show "norm (content k *\<^sub>R (g x - f (m x) x)) \<le> content k * (e / (4 * content (cbox a b)))"
+            unfolding norm_scaleR uv
+            unfolding abs_of_nonneg[OF content_pos_le]
+            apply (rule mult_left_mono)
+            using m(2)[OF x,of "m x"]
+            apply auto
+            done
+        qed (insert False, auto)
+
+      next
+        case 2
+        show ?case
+          apply (rule le_less_trans[of _ "norm (\<Sum>j = 0..s.
+            \<Sum>(x, k)\<in>{xk\<in>p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))"])
+          apply (subst setsum_group)
+          apply fact
+          apply (rule finite_atLeastAtMost)
+          defer
+          apply (subst split_def)+
+          unfolding setsum_subtractf
+          apply rule
+        proof -
+          show "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> p.
+            m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e / 2"
+            apply (rule le_less_trans[of _ "setsum (\<lambda>i. e / 2^(i+2)) {0..s}"])
+            apply (rule setsum_norm_le)
+          proof
+            show "(\<Sum>i = 0..s. e / 2 ^ (i + 2)) < e / 2"
+              unfolding power_add divide_inverse inverse_mult_distrib
+              unfolding setsum_right_distrib[symmetric] setsum_left_distrib[symmetric]
+              unfolding power_inverse [symmetric] sum_gp
+              apply(rule mult_strict_left_mono[OF _ e])
+              unfolding power2_eq_square
+              apply auto
+              done
+            fix t
+            assume "t \<in> {0..s}"
+            show "norm (\<Sum>(x, k)\<in>{xk \<in> p. m (fst xk) = t}. content k *\<^sub>R f (m x) x -
+              integral k (f (m x))) \<le> e / 2 ^ (t + 2)"
+              apply (rule order_trans
+                [of _ "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \<in> p. m (fst xk) = t})"])
+              apply (rule eq_refl)
+              apply (rule arg_cong[where f=norm])
+              apply (rule setsum.cong)
+              apply (rule refl)
+              defer
+              apply (rule henstock_lemma_part1)
+              apply (rule assms(1)[rule_format])
+              apply (simp add: e)
+              apply safe
+              apply (rule c)+
+              apply rule
+              apply assumption+
+              apply (rule tagged_partial_division_subset[of p])
+              apply (rule p(1)[unfolded tagged_division_of_def,THEN conjunct1])
+              defer
+              unfolding fine_def
+              apply safe
+              apply (drule p(2)[unfolded fine_def,rule_format])
+              unfolding d_def
+              apply auto
+              done
+          qed
+        qed (insert s, auto)
+      next
+        case 3
+        note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)]
+        have *: "\<And>sr sx ss ks kr::real. kr = sr \<longrightarrow> ks = ss \<longrightarrow>
+          ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i\<bullet>1 - kr\<bullet>1 \<and> i\<bullet>1 - kr\<bullet>1 < e/4 \<longrightarrow> \<bar>sx - i\<bar> < e/4"
+          by auto
+        show ?case
+          unfolding real_norm_def
+          apply (rule *[rule_format])
+          apply safe
+          apply (rule comb[of r])
+          apply (rule comb[of s])
+          apply (rule i'[unfolded real_inner_1_right])
+          apply (rule_tac[1-2] setsum_mono)
+          unfolding split_paired_all split_conv
+          apply (rule_tac[1-2] integral_le[OF ])
+        proof safe
+          show "0 \<le> i\<bullet>1 - (integral (cbox a b) (f r))\<bullet>1"
+            using r(1) by auto
+          show "i\<bullet>1 - (integral (cbox a b) (f r))\<bullet>1 < e / 4"
+            using r(2) by auto
+          fix x k
+          assume xk: "(x, k) \<in> p"
+          from p'(4)[OF this] guess u v by (elim exE) note uv=this
+          show "f r integrable_on k"
+            and "f s integrable_on k"
+            and "f (m x) integrable_on k"
+            and "f (m x) integrable_on k"
+            unfolding uv
+            apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]])
+            using p'(3)[OF xk]
+            unfolding uv
+            apply auto
+            done
+          fix y
+          assume "y \<in> k"
+          then have "y \<in> cbox a b"
+            using p'(3)[OF xk] by auto
+          then have *: "\<And>m. \<forall>n\<ge>m. f m y \<le> f n y"
+            apply -
+            apply (rule transitive_stepwise_le)
+            using assms(2)
+            apply auto
+            done
+          show "f r y \<le> f (m x) y" and "f (m x) y \<le> f s y"
+            apply (rule_tac[!] *[rule_format])
+            using s[rule_format,OF xk] m(1)[of x] p'(2-3)[OF xk]
+            apply auto
+            done
+        qed
+      qed
+    qed
+  qed note * = this
+
+  have "integral (cbox a b) g = i"
+    by (rule integral_unique) (rule *)
+  then show ?thesis
+    using i * by auto
+qed
+
+lemma monotone_convergence_increasing:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
+  assumes "\<forall>k. (f k) integrable_on s"
+    and "\<forall>k. \<forall>x\<in>s. (f k x) \<le> (f (Suc k) x)"
+    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+    and "bounded {integral s (f k)| k. True}"
+  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+proof -
+  have lem: "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+    if "\<forall>k. \<forall>x\<in>s. 0 \<le> f k x"
+    and "\<forall>k. (f k) integrable_on s"
+    and "\<forall>k. \<forall>x\<in>s. f k x \<le> f (Suc k) x"
+    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+    and "bounded {integral s (f k)| k. True}"
+    for f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" and g s
+  proof -
+    note assms=that[rule_format]
+    have "\<forall>x\<in>s. \<forall>k. (f k x)\<bullet>1 \<le> (g x)\<bullet>1"
+      apply safe
+      apply (rule Lim_component_ge)
+      apply (rule that(4)[rule_format])
+      apply assumption
+      apply (rule trivial_limit_sequentially)
+      unfolding eventually_sequentially
+      apply (rule_tac x=k in exI)
+      apply (rule transitive_stepwise_le)
+      using that(3)
+      apply auto
+      done
+    note fg=this[rule_format]
+
+    have "\<exists>i. ((\<lambda>k. integral s (f k)) \<longlongrightarrow> i) sequentially"
+      apply (rule bounded_increasing_convergent)
+      apply (rule that(5))
+      apply rule
+      apply (rule integral_le)
+      apply (rule that(2)[rule_format])+
+      using that(3)
+      apply auto
+      done
+    then guess i .. note i=this
+    have "\<And>k. \<forall>x\<in>s. \<forall>n\<ge>k. f k x \<le> f n x"
+      apply rule
+      apply (rule transitive_stepwise_le)
+      using that(3)
+      apply auto
+      done
+    then have i': "\<forall>k. (integral s (f k))\<bullet>1 \<le> i\<bullet>1"
+      apply -
+      apply rule
+      apply (rule Lim_component_ge)
+      apply (rule i)
+      apply (rule trivial_limit_sequentially)
+      unfolding eventually_sequentially
+      apply (rule_tac x=k in exI)
+      apply safe
+      apply (rule integral_component_le)
+      apply simp
+      apply (rule that(2)[rule_format])+
+      apply auto
+      done
+
+    note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format]
+    have ifif: "\<And>k t. (\<lambda>x. if x \<in> t then if x \<in> s then f k x else 0 else 0) =
+      (\<lambda>x. if x \<in> t \<inter> s then f k x else 0)"
+      by (rule ext) auto
+    have int': "\<And>k a b. f k integrable_on cbox a b \<inter> s"
+      apply (subst integrable_restrict_univ[symmetric])
+      apply (subst ifif[symmetric])
+      apply (subst integrable_restrict_univ)
+      apply (rule int)
+      done
+    have "\<And>a b. (\<lambda>x. if x \<in> s then g x else 0) integrable_on cbox a b \<and>
+      ((\<lambda>k. integral (cbox a b) (\<lambda>x. if x \<in> s then f k x else 0)) \<longlongrightarrow>
+      integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)) sequentially"
+    proof (rule monotone_convergence_interval, safe, goal_cases)
+      case 1
+      show ?case by (rule int)
+    next
+      case (2 _ _ _ x)
+      then show ?case
+        apply (cases "x \<in> s")
+        using assms(3)
+        apply auto
+        done
+    next
+      case (3 _ _ x)
+      then show ?case
+        apply (cases "x \<in> s")
+        using assms(4)
+        apply auto
+        done
+    next
+      case (4 a b)
+      note * = integral_nonneg
+      have "\<And>k. norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f k x else 0)) \<le> norm (integral s (f k))"
+        unfolding real_norm_def
+        apply (subst abs_of_nonneg)
+        apply (rule *[OF int])
+        apply safe
+        apply (case_tac "x \<in> s")
+        apply (drule assms(1))
+        prefer 3
+        apply (subst abs_of_nonneg)
+        apply (rule *[OF assms(2) that(1)[THEN spec]])
+        apply (subst integral_restrict_univ[symmetric,OF int])
+        unfolding ifif
+        unfolding integral_restrict_univ[OF int']
+        apply (rule integral_subset_le[OF _ int' assms(2)])
+        using assms(1)
+        apply auto
+        done
+      then show ?case
+        using assms(5)
+        unfolding bounded_iff
+        apply safe
+        apply (rule_tac x=aa in exI)
+        apply safe
+        apply (erule_tac x="integral s (f k)" in ballE)
+        apply (rule order_trans)
+        apply assumption
+        apply auto
+        done
+    qed
+    note g = conjunctD2[OF this]
+
+    have "(g has_integral i) s"
+      unfolding has_integral_alt'
+      apply safe
+      apply (rule g(1))
+    proof goal_cases
+      case (1 e)
+      then have "e/4>0"
+        by auto
+      from LIMSEQ_D [OF i this] guess N .. note N=this
+      note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]]
+      from this[THEN conjunct2,rule_format,OF \<open>e/4>0\<close>] guess B .. note B=conjunctD2[OF this]
+      show ?case
+        apply rule
+        apply rule
+        apply (rule B)
+        apply safe
+      proof -
+        fix a b :: 'n
+        assume ab: "ball 0 B \<subseteq> cbox a b"
+        from \<open>e > 0\<close> have "e/2 > 0"
+          by auto
+        from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this
+        have **: "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f N x else 0) - i) < e/2"
+          apply (rule norm_triangle_half_l)
+          using B(2)[rule_format,OF ab] N[rule_format,of N]
+          apply -
+          defer
+          apply (subst norm_minus_commute)
+          apply auto
+          done
+        have *: "\<And>f1 f2 g. \<bar>f1 - i\<bar> < e / 2 \<longrightarrow> \<bar>f2 - g\<bar> < e / 2 \<longrightarrow>
+          f1 \<le> f2 \<longrightarrow> f2 \<le> i \<longrightarrow> \<bar>g - i\<bar> < e"
+          unfolding real_inner_1_right by arith
+        show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0) - i) < e"
+          unfolding real_norm_def
+          apply (rule *[rule_format])
+          apply (rule **[unfolded real_norm_def])
+          apply (rule M[rule_format,of "M + N",unfolded real_norm_def])
+          apply (rule le_add1)
+          apply (rule integral_le[OF int int])
+          defer
+          apply (rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]])
+        proof (safe, goal_cases)
+          case (2 x)
+          have "\<And>m. x \<in> s \<Longrightarrow> \<forall>n\<ge>m. (f m x)\<bullet>1 \<le> (f n x)\<bullet>1"
+            apply (rule transitive_stepwise_le)
+            using assms(3)
+            apply auto
+            done
+          then show ?case
+            by auto
+        next
+          case 1
+          show ?case
+            apply (subst integral_restrict_univ[symmetric,OF int])
+            unfolding ifif integral_restrict_univ[OF int']
+            apply (rule integral_subset_le[OF _ int'])
+            using assms
+            apply auto
+            done
+        qed
+      qed
+    qed
+    then show ?thesis
+      apply safe
+      defer
+      apply (drule integral_unique)
+      using i
+      apply auto
+      done
+  qed
+
+  have sub: "\<And>k. integral s (\<lambda>x. f k x - f 0 x) = integral s (f k) - integral s (f 0)"
+    apply (subst integral_diff)
+    apply (rule assms(1)[rule_format])+
+    apply rule
+    done
+  have "\<And>x m. x \<in> s \<Longrightarrow> \<forall>n\<ge>m. f m x \<le> f n x"
+    apply (rule transitive_stepwise_le)
+    using assms(2)
+    apply auto
+    done
+  note * = this[rule_format]
+  have "(\<lambda>x. g x - f 0 x) integrable_on s \<and> ((\<lambda>k. integral s (\<lambda>x. f (Suc k) x - f 0 x)) \<longlongrightarrow>
+    integral s (\<lambda>x. g x - f 0 x)) sequentially"
+    apply (rule lem)
+    apply safe
+  proof goal_cases
+    case (1 k x)
+    then show ?case
+      using *[of x 0 "Suc k"] by auto
+  next
+    case (2 k)
+    then show ?case
+      apply (rule integrable_diff)
+      using assms(1)
+      apply auto
+      done
+  next
+    case (3 k x)
+    then show ?case
+      using *[of x "Suc k" "Suc (Suc k)"] by auto
+  next
+    case (4 x)
+    then show ?case
+      apply -
+      apply (rule tendsto_diff)
+      using LIMSEQ_ignore_initial_segment[OF assms(3)[rule_format],of x 1]
+      apply auto
+      done
+  next
+    case 5
+    then show ?case
+      using assms(4)
+      unfolding bounded_iff
+      apply safe
+      apply (rule_tac x="a + norm (integral s (\<lambda>x. f 0 x))" in exI)
+      apply safe
+      apply (erule_tac x="integral s (\<lambda>x. f (Suc k) x)" in ballE)
+      unfolding sub
+      apply (rule order_trans[OF norm_triangle_ineq4])
+      apply auto
+      done
+  qed
+  note conjunctD2[OF this]
+  note tendsto_add[OF this(2) tendsto_const[of "integral s (f 0)"]]
+    integrable_add[OF this(1) assms(1)[rule_format,of 0]]
+  then show ?thesis
+    unfolding sub
+    apply -
+    apply rule
+    defer
+    apply (subst(asm) integral_diff)
+    using assms(1)
+    apply auto
+    apply (rule LIMSEQ_imp_Suc)
+    apply assumption
+    done
+qed
+
+lemma has_integral_monotone_convergence_increasing:
+  fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow> real"
+  assumes f: "\<And>k. (f k has_integral x k) s"
+  assumes "\<And>k x. x \<in> s \<Longrightarrow> f k x \<le> f (Suc k) x"
+  assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
+  assumes "x \<longlonglongrightarrow> x'"
+  shows "(g has_integral x') s"
+proof -
+  have x_eq: "x = (\<lambda>i. integral s (f i))"
+    by (simp add: integral_unique[OF f])
+  then have x: "{integral s (f k) |k. True} = range x"
+    by auto
+
+  have *: "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
+  proof (intro monotone_convergence_increasing allI ballI assms)
+    show "bounded {integral s (f k) |k. True}"
+      unfolding x by (rule convergent_imp_bounded) fact
+  qed (auto intro: f)
+  then have "integral s g = x'"
+    by (intro LIMSEQ_unique[OF _ \<open>x \<longlonglongrightarrow> x'\<close>]) (simp add: x_eq)
+  with * show ?thesis
+    by (simp add: has_integral_integral)
+qed
+
+lemma monotone_convergence_decreasing:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
+  assumes "\<forall>k. (f k) integrable_on s"
+    and "\<forall>k. \<forall>x\<in>s. f (Suc k) x \<le> f k x"
+    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+    and "bounded {integral s (f k)| k. True}"
+  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+proof -
+  note assm = assms[rule_format]
+  have *: "{integral s (\<lambda>x. - f k x) |k. True} = op *\<^sub>R (- 1) ` {integral s (f k)| k. True}"
+    apply safe
+    unfolding image_iff
+    apply (rule_tac x="integral s (f k)" in bexI)
+    prefer 3
+    apply (rule_tac x=k in exI)
+    apply auto
+    done
+  have "(\<lambda>x. - g x) integrable_on s \<and>
+    ((\<lambda>k. integral s (\<lambda>x. - f k x)) \<longlongrightarrow> integral s (\<lambda>x. - g x)) sequentially"
+    apply (rule monotone_convergence_increasing)
+    apply safe
+    apply (rule integrable_neg)
+    apply (rule assm)
+    defer
+    apply (rule tendsto_minus)
+    apply (rule assm)
+    apply assumption
+    unfolding *
+    apply (rule bounded_scaling)
+    using assm
+    apply auto
+    done
+  note * = conjunctD2[OF this]
+  show ?thesis
+    using integrable_neg[OF *(1)] tendsto_minus[OF *(2)]
+    by auto
+qed
+
+
+subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close>
+
+definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46)
+  where "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and> (\<lambda>x. (norm(f x))) integrable_on s"
+
+lemma absolutely_integrable_onI[intro?]:
+  "f integrable_on s \<Longrightarrow>
+    (\<lambda>x. (norm(f x))) integrable_on s \<Longrightarrow> f absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def
+  by auto
+
+lemma absolutely_integrable_onD[dest]:
+  assumes "f absolutely_integrable_on s"
+  shows "f integrable_on s"
+    and "(\<lambda>x. norm (f x)) integrable_on s"
+  using assms
+  unfolding absolutely_integrable_on_def
+  by auto
+
+(*lemma absolutely_integrable_on_trans[simp]:
+  fixes f::"'n::euclidean_space \<Rightarrow> real"
+  shows "(vec1 \<circ> f) absolutely_integrable_on s \<longleftrightarrow> f absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def o_def by auto*)
+
+lemma integral_norm_bound_integral:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f integrable_on s"
+    and "g integrable_on s"
+    and "\<forall>x\<in>s. norm (f x) \<le> g x"
+  shows "norm (integral s f) \<le> integral s g"
+proof -
+  have *: "\<And>x y. (\<forall>e::real. 0 < e \<longrightarrow> x < y + e) \<Longrightarrow> x \<le> y"
+    apply (rule ccontr)
+    apply (erule_tac x="x - y" in allE)
+    apply auto
+    done
+  have norm: "norm ig < dia + e"
+    if "norm sg \<le> dsa"
+    and "\<bar>dsa - dia\<bar> < e / 2"
+    and "norm (sg - ig) < e / 2"
+    for e dsa dia and sg ig :: 'a
+    apply (rule le_less_trans[OF norm_triangle_sub[of ig sg]])
+    apply (subst real_sum_of_halves[of e,symmetric])
+    unfolding add.assoc[symmetric]
+    apply (rule add_le_less_mono)
+    defer
+    apply (subst norm_minus_commute)
+    apply (rule that(3))
+    apply (rule order_trans[OF that(1)])
+    using that(2)
+    apply arith
+    done
+  have lem: "norm (integral(cbox a b) f) \<le> integral (cbox a b) g"
+    if "f integrable_on cbox a b"
+    and "g integrable_on cbox a b"
+    and "\<forall>x\<in>cbox a b. norm (f x) \<le> g x"
+    for f :: "'n \<Rightarrow> 'a" and g a b
+  proof (rule *[rule_format])
+    fix e :: real
+    assume "e > 0"
+    then have *: "e/2 > 0"
+      by auto
+    from integrable_integral[OF that(1),unfolded has_integral[of f],rule_format,OF *]
+    guess d1 .. note d1 = conjunctD2[OF this,rule_format]
+    from integrable_integral[OF that(2),unfolded has_integral[of g],rule_format,OF *]
+    guess d2 .. note d2 = conjunctD2[OF this,rule_format]
+    note gauge_inter[OF d1(1) d2(1)]
+    from fine_division_exists[OF this, of a b] guess p . note p=this
+    show "norm (integral (cbox a b) f) < integral (cbox a b) g + e"
+      apply (rule norm)
+      defer
+      apply (rule d2(2)[OF conjI[OF p(1)],unfolded real_norm_def])
+      defer
+      apply (rule d1(2)[OF conjI[OF p(1)]])
+      defer
+      apply (rule setsum_norm_le)
+    proof safe
+      fix x k
+      assume "(x, k) \<in> p"
+      note as = tagged_division_ofD(2-4)[OF p(1) this]
+      from this(3) guess u v by (elim exE) note uv=this
+      show "norm (content k *\<^sub>R f x) \<le> content k *\<^sub>R g x"
+        unfolding uv norm_scaleR
+        unfolding abs_of_nonneg[OF content_pos_le] real_scaleR_def
+        apply (rule mult_left_mono)
+        using that(3) as
+        apply auto
+        done
+    qed (insert p[unfolded fine_inter], auto)
+  qed
+
+  { presume "\<And>e. 0 < e \<Longrightarrow> norm (integral s f) < integral s g + e"
+    then show ?thesis by (rule *[rule_format]) auto }
+  fix e :: real
+  assume "e > 0"
+  then have e: "e/2 > 0"
+    by auto
+  note assms(1)[unfolded integrable_alt[of f]] note f=this[THEN conjunct1,rule_format]
+  note assms(2)[unfolded integrable_alt[of g]] note g=this[THEN conjunct1,rule_format]
+  from integrable_integral[OF assms(1),unfolded has_integral'[of f],rule_format,OF e]
+  guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format]
+  from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e]
+  guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format]
+  from bounded_subset_cbox[OF bounded_ball, of "0::'n" "max B1 B2"]
+  guess a b by (elim exE) note ab=this[unfolded ball_max_Un]
+
+  have "ball 0 B1 \<subseteq> cbox a b"
+    using ab by auto
+  from B1(2)[OF this] guess z .. note z=conjunctD2[OF this]
+  have "ball 0 B2 \<subseteq> cbox a b"
+    using ab by auto
+  from B2(2)[OF this] guess w .. note w=conjunctD2[OF this]
+
+  show "norm (integral s f) < integral s g + e"
+    apply (rule norm)
+    apply (rule lem[OF f g, of a b])
+    unfolding integral_unique[OF z(1)] integral_unique[OF w(1)]
+    defer
+    apply (rule w(2)[unfolded real_norm_def])
+    apply (rule z(2))
+    apply safe
+    apply (case_tac "x \<in> s")
+    unfolding if_P
+    apply (rule assms(3)[rule_format])
+    apply auto
+    done
+qed
+
+lemma integral_norm_bound_integral_component:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
+  assumes "f integrable_on s"
+    and "g integrable_on s"
+    and "\<forall>x\<in>s. norm(f x) \<le> (g x)\<bullet>k"
+  shows "norm (integral s f) \<le> (integral s g)\<bullet>k"
+proof -
+  have "norm (integral s f) \<le> integral s ((\<lambda>x. x \<bullet> k) \<circ> g)"
+    apply (rule integral_norm_bound_integral[OF assms(1)])
+    apply (rule integrable_linear[OF assms(2)])
+    apply rule
+    unfolding o_def
+    apply (rule assms)
+    done
+  then show ?thesis
+    unfolding o_def integral_component_eq[OF assms(2)] .
+qed
+
+lemma has_integral_norm_bound_integral_component:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
+  assumes "(f has_integral i) s"
+    and "(g has_integral j) s"
+    and "\<forall>x\<in>s. norm (f x) \<le> (g x)\<bullet>k"
+  shows "norm i \<le> j\<bullet>k"
+  using integral_norm_bound_integral_component[of f s g k]
+  unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)]
+  using assms
+  by auto
+
+lemma absolutely_integrable_le:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f absolutely_integrable_on s"
+  shows "norm (integral s f) \<le> integral s (\<lambda>x. norm (f x))"
+  apply (rule integral_norm_bound_integral)
+  using assms
+  apply auto
+  done
+
+lemma absolutely_integrable_0[intro]:
+  "(\<lambda>x. 0) absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def
+  by auto
+
+lemma absolutely_integrable_cmul[intro]:
+  "f absolutely_integrable_on s \<Longrightarrow>
+    (\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def
+  using integrable_cmul[of f s c]
+  using integrable_cmul[of "\<lambda>x. norm (f x)" s "\<bar>c\<bar>"]
+  by auto
+
+lemma absolutely_integrable_neg[intro]:
+  "f absolutely_integrable_on s \<Longrightarrow>
+    (\<lambda>x. -f(x)) absolutely_integrable_on s"
+  apply (drule absolutely_integrable_cmul[where c="-1"])
+  apply auto
+  done
+
+lemma absolutely_integrable_norm[intro]:
+  "f absolutely_integrable_on s \<Longrightarrow>
+    (\<lambda>x. norm (f x)) absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def
+  by auto
+
+lemma absolutely_integrable_abs[intro]:
+  "f absolutely_integrable_on s \<Longrightarrow>
+    (\<lambda>x. \<bar>f x::real\<bar>) absolutely_integrable_on s"
+  apply (drule absolutely_integrable_norm)
+  unfolding real_norm_def
+  apply assumption
+  done
+
+lemma absolutely_integrable_on_subinterval:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  shows "f absolutely_integrable_on s \<Longrightarrow>
+    cbox a b \<subseteq> s \<Longrightarrow> f absolutely_integrable_on cbox a b"
+  unfolding absolutely_integrable_on_def
+  by (metis integrable_on_subcbox)
+
+lemma absolutely_integrable_bounded_variation:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
+  assumes "f absolutely_integrable_on UNIV"
+  obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
+  apply (rule that[of "integral UNIV (\<lambda>x. norm (f x))"])
+  apply safe
+proof goal_cases
+  case prems: (1 d)
+  note d = division_ofD[OF prems(2)]
+  have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. integral i (\<lambda>x. norm (f x)))"
+    apply (rule setsum_mono,rule absolutely_integrable_le)
+    apply (drule d(4))
+    apply safe
+    apply (rule absolutely_integrable_on_subinterval[OF assms])
+    apply auto
+    done
+  also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))"
+    apply (subst integral_combine_division_topdown[OF _ prems(2)])
+    using integrable_on_subdivision[OF prems(2)]
+    using assms
+    apply auto
+    done
+  also have "\<dots> \<le> integral UNIV (\<lambda>x. norm (f x))"
+    apply (rule integral_subset_le)
+    using integrable_on_subdivision[OF prems(2)]
+    using assms
+    apply auto
+    done
+  finally show ?case .
+qed
+
+lemma helplemma:
+  assumes "setsum (\<lambda>x. norm (f x - g x)) s < e"
+    and "finite s"
+  shows "\<bar>setsum (\<lambda>x. norm(f x)) s - setsum (\<lambda>x. norm(g x)) s\<bar> < e"
+  unfolding setsum_subtractf[symmetric]
+  apply (rule le_less_trans[OF setsum_abs])
+  apply (rule le_less_trans[OF _ assms(1)])
+  apply (rule setsum_mono)
+  apply (rule norm_triangle_ineq3)
+  done
+
+lemma bounded_variation_absolutely_integrable_interval:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes f: "f integrable_on cbox a b"
+    and *: "\<forall>d. d division_of (cbox a b) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
+  shows "f absolutely_integrable_on cbox a b"
+proof -
+  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}"
+  have D_1: "?D \<noteq> {}"
+    by (rule elementary_interval[of a b]) auto
+  have D_2: "bdd_above (?f`?D)"
+    by (metis * mem_Collect_eq bdd_aboveI2)
+  note D = D_1 D_2
+  let ?S = "SUP x:?D. ?f x"
+  show ?thesis
+    apply (rule absolutely_integrable_onI [OF f has_integral_integrable])
+    apply (subst has_integral[of _ ?S])
+    apply safe
+  proof goal_cases
+    case e: (1 e)
+    then have "?S - e / 2 < ?S" by simp
+    then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\<Sum>k\<in>d. norm (integral k f))"
+      unfolding less_cSUP_iff[OF D] by auto
+    note d' = division_ofD[OF this(1)]
+
+    have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
+    proof
+      fix x
+      have "\<exists>da>0. \<forall>xa\<in>\<Union>{i \<in> d. x \<notin> i}. da \<le> dist x xa"
+        apply (rule separate_point_closed)
+        apply (rule closed_Union)
+        apply (rule finite_subset[OF _ d'(1)])
+        using d'(4)
+        apply auto
+        done
+      then show "\<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
+        by force
+    qed
+    from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format]
+
+    have "e/2 > 0"
+      using e by auto
+    from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format]
+    let ?g = "\<lambda>x. g x \<inter> ball x (k x)"
+    show ?case
+      apply (rule_tac x="?g" in exI)
+      apply safe
+    proof -
+      show "gauge ?g"
+        using g(1) k(1)
+        unfolding gauge_def
+        by auto
+      fix p
+      assume "p tagged_division_of (cbox a b)" and "?g fine p"
+      note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]]
+      note p' = tagged_division_ofD[OF p(1)]
+      define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
+      have gp': "g fine p'"
+        using p(2)
+        unfolding p'_def fine_def
+        by auto
+      have p'': "p' tagged_division_of (cbox a b)"
+        apply (rule tagged_division_ofI)
+      proof -
+        show "finite p'"
+          apply (rule finite_subset[of _ "(\<lambda>(k,(x,l)). (x,k \<inter> l)) `
+            {(k,xl) | k xl. k \<in> d \<and> xl \<in> p}"])
+          unfolding p'_def
+          defer
+          apply (rule finite_imageI,rule finite_product_dependent[OF d'(1) p'(1)])
+          apply safe
+          unfolding image_iff
+          apply (rule_tac x="(i,x,l)" in bexI)
+          apply auto
+          done
+        fix x k
+        assume "(x, k) \<in> p'"
+        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
+          unfolding p'_def by auto
+        then guess i l by (elim exE) note il=conjunctD4[OF this]
+        show "x \<in> k" and "k \<subseteq> cbox a b"
+          using p'(2-3)[OF il(3)] il by auto
+        show "\<exists>a b. k = cbox a b"
+          unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)]
+          apply safe
+          unfolding inter_interval
+          apply auto
+          done
+      next
+        fix x1 k1
+        assume "(x1, k1) \<in> p'"
+        then have "\<exists>i l. x1 \<in> i \<and> i \<in> d \<and> (x1, l) \<in> p \<and> k1 = i \<inter> l"
+          unfolding p'_def by auto
+        then guess i1 l1 by (elim exE) note il1=conjunctD4[OF this]
+        fix x2 k2
+        assume "(x2,k2)\<in>p'"
+        then have "\<exists>i l. x2 \<in> i \<and> i \<in> d \<and> (x2, l) \<in> p \<and> k2 = i \<inter> l"
+          unfolding p'_def by auto
+        then guess i2 l2 by (elim exE) note il2=conjunctD4[OF this]
+        assume "(x1, k1) \<noteq> (x2, k2)"
+        then have "interior i1 \<inter> interior i2 = {} \<or> interior l1 \<inter> interior l2 = {}"
+          using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)]
+          unfolding il1 il2
+          by auto
+        then show "interior k1 \<inter> interior k2 = {}"
+          unfolding il1 il2 by auto
+      next
+        have *: "\<forall>(x, X) \<in> p'. X \<subseteq> cbox a b"
+          unfolding p'_def using d' by auto
+        show "\<Union>{k. \<exists>x. (x, k) \<in> p'} = cbox a b"
+          apply rule
+          apply (rule Union_least)
+          unfolding mem_Collect_eq
+          apply (erule exE)
+          apply (drule *[rule_format])
+          apply safe
+        proof -
+          fix y
+          assume y: "y \<in> cbox a b"
+          then have "\<exists>x l. (x, l) \<in> p \<and> y\<in>l"
+            unfolding p'(6)[symmetric] by auto
+          then guess x l by (elim exE) note xl=conjunctD2[OF this]
+          then have "\<exists>k. k \<in> d \<and> y \<in> k"
+            using y unfolding d'(6)[symmetric] by auto
+          then guess i .. note i = conjunctD2[OF this]
+          have "x \<in> i"
+            using fineD[OF p(3) xl(1)]
+            using k(2)[OF i(1), of x]
+            using i(2) xl(2)
+            by auto
+          then show "y \<in> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
+            unfolding p'_def Union_iff
+            apply (rule_tac x="i \<inter> l" in bexI)
+            using i xl
+            apply auto
+            done
+        qed
+      qed
+
+      then have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2"
+        apply -
+        apply (rule g(2)[rule_format])
+        unfolding tagged_division_of_def
+        apply safe
+        apply (rule gp')
+        done
+      then have **: "\<bar>(\<Sum>(x,k)\<in>p'. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p'. norm (integral k f))\<bar> < e / 2"
+        unfolding split_def
+        using p''
+        by (force intro!: helplemma)
+
+      have p'alt: "p' = {(x,(i \<inter> l)) | x i l. (x,l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}}"
+      proof (safe, goal_cases)
+        case prems: (2 _ _ x i l)
+        have "x \<in> i"
+          using fineD[OF p(3) prems(1)] k(2)[OF prems(2), of x] prems(4-)
+          by auto
+        then have "(x, i \<inter> l) \<in> p'"
+          unfolding p'_def
+          using prems
+          apply safe
+          apply (rule_tac x=x in exI)
+          apply (rule_tac x="i \<inter> l" in exI)
+          apply safe
+          using prems
+          apply auto
+          done
+        then show ?case
+          using prems(3) by auto
+      next
+        fix x k
+        assume "(x, k) \<in> p'"
+        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
+          unfolding p'_def by auto
+        then guess i l by (elim exE) note il=conjunctD4[OF this]
+        then show "\<exists>y i l. (x, k) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
+          apply (rule_tac x=x in exI)
+          apply (rule_tac x=i in exI)
+          apply (rule_tac x=l in exI)
+          using p'(2)[OF il(3)]
+          apply auto
+          done
+      qed
+      have sum_p': "(\<Sum>(x, k)\<in>p'. norm (integral k f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
+        apply (subst setsum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
+        unfolding norm_eq_zero
+        apply (rule integral_null)
+        apply assumption
+        apply rule
+        done
+      note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
+
+      have *: "\<And>sni sni' sf sf'. \<bar>sf' - sni'\<bar> < e / 2 \<longrightarrow> ?S - e / 2 < sni \<and> sni' \<le> ?S \<and>
+        sni \<le> sni' \<and> sf' = sf \<longrightarrow> \<bar>sf - ?S\<bar> < e"
+        by arith
+      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - ?S) < e"
+        unfolding real_norm_def
+        apply (rule *[rule_format,OF **])
+        apply safe
+        apply(rule d(2))
+      proof goal_cases
+        case 1
+        show ?case
+          by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
+      next
+        case 2
+        have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} =
+          (\<lambda>(k,l). k \<inter> l) ` {(k,l)|k l. k \<in> d \<and> l \<in> snd ` p}"
+          by auto
+        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))"
+        proof (rule setsum_mono, goal_cases)
+          case k: (1 k)
+          from d'(4)[OF this] guess u v by (elim exE) note uv=this
+          define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and>  cbox u v \<inter> l \<noteq> {}}"
+          note uvab = d'(2)[OF k[unfolded uv]]
+          have "d' division_of cbox u v"
+            apply (subst d'_def)
+            apply (rule division_inter_1)
+            apply (rule division_of_tagged_division[OF p(1)])
+            apply (rule uvab)
+            done
+          then have "norm (integral k f) \<le> setsum (\<lambda>k. norm (integral k f)) d'"
+            unfolding uv
+            apply (subst integral_combine_division_topdown[of _ _ d'])
+            apply (rule integrable_on_subcbox[OF assms(1) uvab])
+            apply assumption
+            apply (rule setsum_norm_le)
+            apply auto
+            done
+          also have "\<dots> = (\<Sum>k\<in>{k \<inter> l |l. l \<in> snd ` p}. norm (integral k f))"
+            apply (rule setsum.mono_neutral_left)
+            apply (subst simple_image)
+            apply (rule finite_imageI)+
+            apply fact
+            unfolding d'_def uv
+            apply blast
+          proof (rule, goal_cases)
+            case prems: (1 i)
+            then have "i \<in> {cbox u v \<inter> l |l. l \<in> snd ` p}"
+              by auto
+            from this[unfolded mem_Collect_eq] guess l .. note l=this
+            then have "cbox u v \<inter> l = {}"
+              using prems by auto
+            then show ?case
+              using l by auto
+          qed
+          also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (k \<inter> l) f))"
+            unfolding simple_image
+            apply (rule setsum.reindex_nontrivial [unfolded o_def])
+            apply (rule finite_imageI)
+            apply (rule p')
+          proof goal_cases
+            case prems: (1 l y)
+            have "interior (k \<inter> l) \<subseteq> interior (l \<inter> y)"
+              apply (subst(2) interior_Int)
+              apply (rule Int_greatest)
+              defer
+              apply (subst prems(4))
+              apply auto
+              done
+            then have *: "interior (k \<inter> l) = {}"
+              using snd_p(5)[OF prems(1-3)] by auto
+            from d'(4)[OF k] snd_p(4)[OF prems(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
+            show ?case
+              using *
+              unfolding uv inter_interval content_eq_0_interior[symmetric]
+              by auto
+          qed
+          finally show ?case .
+        qed
+        also have "\<dots> = (\<Sum>(i,l)\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (i\<inter>l) f))"
+          apply (subst sum_sum_product[symmetric])
+          apply fact
+          using p'(1)
+          apply auto
+          done
+        also have "\<dots> = (\<Sum>x\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (case_prod op \<inter> x) f))"
+          unfolding split_def ..
+        also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
+          unfolding *
+          apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def])
+          apply (rule finite_product_dependent)
+          apply fact
+          apply (rule finite_imageI)
+          apply (rule p')
+          unfolding split_paired_all mem_Collect_eq split_conv o_def
+        proof -
+          note * = division_ofD(4,5)[OF division_of_tagged_division,OF p(1)]
+          fix l1 l2 k1 k2
+          assume as:
+            "(l1, k1) \<noteq> (l2, k2)"
+            "l1 \<inter> k1 = l2 \<inter> k2"
+            "\<exists>i l. (l1, k1) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
+            "\<exists>i l. (l2, k2) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
+          then have "l1 \<in> d" and "k1 \<in> snd ` p"
+            by auto from d'(4)[OF this(1)] *(1)[OF this(2)]
+          guess u1 v1 u2 v2 by (elim exE) note uv=this
+          have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
+            using as by auto
+          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
+            apply -
+            apply (erule disjE)
+            apply (rule disjI2)
+            apply (rule d'(5))
+            prefer 4
+            apply (rule disjI1)
+            apply (rule *)
+            using as
+            apply auto
+            done
+          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
+            using as(2) by auto
+          ultimately have "interior(l1 \<inter> k1) = {}"
+            by auto
+          then show "norm (integral (l1 \<inter> k1) f) = 0"
+            unfolding uv inter_interval
+            unfolding content_eq_0_interior[symmetric]
+            by auto
+        qed
+        also have "\<dots> = (\<Sum>(x, k)\<in>p'. norm (integral k f))"
+          unfolding sum_p'
+          apply (rule setsum.mono_neutral_right)
+          apply (subst *)
+          apply (rule finite_imageI[OF finite_product_dependent])
+          apply fact
+          apply (rule finite_imageI[OF p'(1)])
+          apply safe
+        proof goal_cases
+          case (2 i ia l a b)
+          then have "ia \<inter> b = {}"
+            unfolding p'alt image_iff Bex_def not_ex
+            apply (erule_tac x="(a, ia \<inter> b)" in allE)
+            apply auto
+            done
+          then show ?case
+            by auto
+        next
+          case (1 x a b)
+          then show ?case
+            unfolding p'_def
+            apply safe
+            apply (rule_tac x=i in exI)
+            apply (rule_tac x=l in exI)
+            unfolding snd_conv image_iff
+            apply safe
+            apply (rule_tac x="(a,l)" in bexI)
+            apply auto
+            done
+        qed
+        finally show ?case .
+      next
+        case 3
+        let ?S = "{(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
+        have Sigma_alt: "\<And>s t. s \<times> t = {(i, j) |i j. i \<in> s \<and> j \<in> t}"
+          by auto
+        have *: "?S = (\<lambda>(xl,i). (fst xl, snd xl \<inter> i)) ` (p \<times> d)"
+          apply safe
+          unfolding image_iff
+          apply (rule_tac x="((x,l),i)" in bexI)
+          apply auto
+          done
+        note pdfin = finite_cartesian_product[OF p'(1) d'(1)]
+        have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x, k)\<in>?S. \<bar>content k\<bar> * norm (f x))"
+          unfolding norm_scaleR
+          apply (rule setsum.mono_neutral_left)
+          apply (subst *)
+          apply (rule finite_imageI)
+          apply fact
+          unfolding p'alt
+          apply blast
+          apply safe
+          apply (rule_tac x=x in exI)
+          apply (rule_tac x=i in exI)
+          apply (rule_tac x=l in exI)
+          apply auto
+          done
+        also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
+          unfolding *
+          apply (subst setsum.reindex_nontrivial)
+          apply fact
+          unfolding split_paired_all
+          unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject
+          apply (elim conjE)
+        proof -
+          fix x1 l1 k1 x2 l2 k2
+          assume as: "(x1, l1) \<in> p" "(x2, l2) \<in> p" "k1 \<in> d" "k2 \<in> d"
+            "x1 = x2" "l1 \<inter> k1 = l2 \<inter> k2" "\<not> ((x1 = x2 \<and> l1 = l2) \<and> k1 = k2)"
+          from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
+          from as have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
+            by auto
+          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
+            apply -
+            apply (erule disjE)
+            apply (rule disjI2)
+            defer
+            apply (rule disjI1)
+            apply (rule d'(5)[OF as(3-4)])
+            apply assumption
+            apply (rule p'(5)[OF as(1-2)])
+            apply auto
+            done
+          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
+            unfolding  as ..
+          ultimately have "interior (l1 \<inter> k1) = {}"
+            by auto
+          then show "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0"
+            unfolding uv inter_interval
+            unfolding content_eq_0_interior[symmetric]
+            by auto
+        qed safe
+        also have "\<dots> = (\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x))"
+          unfolding Sigma_alt
+          apply (subst sum_sum_product[symmetric])
+          apply (rule p')
+          apply rule
+          apply (rule d')
+          apply (rule setsum.cong)
+          apply (rule refl)
+          unfolding split_paired_all split_conv
+        proof -
+          fix x l
+          assume as: "(x, l) \<in> p"
+          note xl = p'(2-4)[OF this]
+          from this(3) guess u v by (elim exE) note uv=this
+          have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
+            apply (rule setsum.cong)
+            apply (rule refl)
+            apply (drule d'(4))
+            apply safe
+            apply (subst Int_commute)
+            unfolding inter_interval uv
+            apply (subst abs_of_nonneg)
+            apply auto
+            done
+          also have "\<dots> = setsum content {k \<inter> cbox u v| k. k \<in> d}"
+            unfolding simple_image
+            apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric])
+            apply (rule d')
+          proof goal_cases
+            case prems: (1 k y)
+            from d'(4)[OF this(1)] d'(4)[OF this(2)]
+            guess u1 v1 u2 v2 by (elim exE) note uv=this
+            have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
+              apply (subst interior_Int)
+              using d'(5)[OF prems(1-3)]
+              apply auto
+              done
+            also have "\<dots> = interior (y \<inter> (k \<inter> cbox u v))"
+              by auto
+            also have "\<dots> = interior (k \<inter> cbox u v)"
+              unfolding prems(4) by auto
+            finally show ?case
+              unfolding uv inter_interval content_eq_0_interior ..
+          qed
+          also have "\<dots> = setsum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
+            apply (rule setsum.mono_neutral_right)
+            unfolding simple_image
+            apply (rule finite_imageI)
+            apply (rule d')
+            apply blast
+            apply safe
+            apply (rule_tac x=k in exI)
+          proof goal_cases
+            case prems: (1 i k)
+            from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this
+            have "interior (k \<inter> cbox u v) \<noteq> {}"
+              using prems(2)
+              unfolding ab inter_interval content_eq_0_interior
+              by auto
+            then show ?case
+              using prems(1)
+              using interior_subset[of "k \<inter> cbox u v"]
+              by auto
+          qed
+          finally show "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar> * norm (f x)) = content l *\<^sub>R norm (f x)"
+            unfolding setsum_left_distrib[symmetric] real_scaleR_def
+            apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
+            using xl(2)[unfolded uv]
+            unfolding uv
+            apply auto
+            done
+        qed
+        finally show ?case .
+      qed
+    qed
+  qed
+qed
+
+lemma bounded_variation_absolutely_integrable:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "f integrable_on UNIV"
+    and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
+  shows "f absolutely_integrable_on UNIV"
+proof (rule absolutely_integrable_onI, fact, rule)
+  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
+  have D_1: "?D \<noteq> {}"
+    by (rule elementary_interval) auto
+  have D_2: "bdd_above (?f`?D)"
+    by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
+  note D = D_1 D_2
+  let ?S = "SUP d:?D. ?f d"
+  have f_int: "\<And>a b. f absolutely_integrable_on cbox a b"
+    apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
+    apply (rule integrable_on_subcbox[OF assms(1)])
+    defer
+    apply safe
+    apply (rule assms(2)[rule_format])
+    apply auto
+    done
+  show "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
+    apply (subst has_integral_alt')
+    apply safe
+  proof goal_cases
+    case (1 a b)
+    show ?case
+      using f_int[of a b] by auto
+  next
+    case prems: (2 e)
+    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
+    proof (rule ccontr)
+      assume "\<not> ?thesis"
+      then have "?S \<le> ?S - e"
+        by (intro cSUP_least[OF D(1)]) auto
+      then show False
+        using prems by auto
+    qed
+    then obtain K where *: "\<exists>x\<in>{d. d division_of \<Union>d}. K = (\<Sum>k\<in>x. norm (integral k f))"
+      "SUPREMUM {d. d division_of \<Union>d} (setsum (\<lambda>k. norm (integral k f))) - e < K"
+      by (auto simp add: image_iff not_le)
+    from this(1) obtain d where "d division_of \<Union>d"
+      and "K = (\<Sum>k\<in>d. norm (integral k f))"
+      by auto
+    note d = this(1) *(2)[unfolded this(2)]
+    note d'=division_ofD[OF this(1)]
+    have "bounded (\<Union>d)"
+      by (rule elementary_bounded,fact)
+    from this[unfolded bounded_pos] obtain K where
+       K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" by auto
+    show ?case
+      apply (rule_tac x="K + 1" in exI)
+      apply safe
+    proof -
+      fix a b :: 'n
+      assume ab: "ball 0 (K + 1) \<subseteq> cbox a b"
+      have *: "\<forall>s s1. ?S - e < s1 \<and> s1 \<le> s \<and> s < ?S + e \<longrightarrow> \<bar>s - ?S\<bar> < e"
+        by arith
+      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
+        unfolding real_norm_def
+        apply (rule *[rule_format])
+        apply safe
+        apply (rule d(2))
+      proof goal_cases
+        case 1
+        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> setsum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
+          apply (rule setsum_mono)
+          apply (rule absolutely_integrable_le)
+          apply (drule d'(4))
+          apply safe
+          apply (rule f_int)
+          done
+        also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))"
+          apply (rule integral_combine_division_bottomup[symmetric])
+          apply (rule d)
+          unfolding forall_in_division[OF d(1)]
+          using f_int
+          apply auto
+          done
+        also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)"
+        proof -
+          have "\<Union>d \<subseteq> cbox a b"
+            apply rule
+            apply (drule K(2)[rule_format])
+            apply (rule ab[unfolded subset_eq,rule_format])
+            apply (auto simp add: dist_norm)
+            done
+          then show ?thesis
+            apply -
+            apply (subst if_P)
+            apply rule
+            apply (rule integral_subset_le)
+            defer
+            apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"])
+            apply (rule d)
+            using f_int[of a b]
+            apply auto
+            done
+        qed
+        finally show ?case .
+      next
+        note f = absolutely_integrable_onD[OF f_int[of a b]]
+        note * = this(2)[unfolded has_integral_integral has_integral[of "\<lambda>x. norm (f x)"],rule_format]
+        have "e/2>0"
+          using \<open>e > 0\<close> by auto
+        from * [OF this] obtain d1 where
+          d1: "gauge d1" "\<forall>p. p tagged_division_of (cbox a b) \<and> d1 fine p \<longrightarrow>
+            norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e / 2"
+          by auto
+        from henstock_lemma [OF f(1) \<open>e/2>0\<close>] obtain d2 where
+          d2: "gauge d2" "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d2 fine p \<longrightarrow>
+            (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" .
+         obtain p where
+          p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p"
+          by (rule fine_division_exists [OF gauge_inter [OF d1(1) d2(1)], of a b])
+            (auto simp add: fine_inter)
+        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> ?S \<longrightarrow> \<bar>sf - si\<bar> < e / 2 \<longrightarrow>
+          \<bar>sf' - di\<bar> < e / 2 \<longrightarrow> di < ?S + e"
+          by arith
+        show "integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
+          apply (subst if_P)
+          apply rule
+        proof (rule *[rule_format])
+          show "\<bar>(\<Sum>(x,k)\<in>p. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p. norm (integral k f))\<bar> < e / 2"
+            unfolding split_def
+            apply (rule helplemma)
+            using d2(2)[rule_format,of p]
+            using p(1,3)
+            unfolding tagged_division_of_def split_def
+            apply auto
+            done
+          show "\<bar>(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm(f x))\<bar> < e / 2"
+            using d1(2)[rule_format,OF conjI[OF p(1,2)]]
+            by (simp only: real_norm_def)
+          show "(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) = (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x))"
+            apply (rule setsum.cong)
+            apply (rule refl)
+            unfolding split_paired_all split_conv
+            apply (drule tagged_division_ofD(4)[OF p(1)])
+            unfolding norm_scaleR
+            apply (subst abs_of_nonneg)
+            apply auto
+            done
+          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
+            using partial_division_of_tagged_division[of p "cbox a b"] p(1)
+            apply (subst setsum.over_tagged_division_lemma[OF p(1)])
+            apply (simp add: integral_null)
+            apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
+            apply (auto simp: tagged_partial_division_of_def)
+            done
+        qed
+      qed
+    qed (insert K, auto)
+  qed
+qed
+
+lemma absolutely_integrable_restrict_univ:
+  "(\<lambda>x. if x \<in> s then f x else (0::'a::banach)) absolutely_integrable_on UNIV \<longleftrightarrow>
+    f absolutely_integrable_on s"
+  unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ ..
+
+lemma absolutely_integrable_add[intro]:
+  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "f absolutely_integrable_on s"
+    and "g absolutely_integrable_on s"
+  shows "(\<lambda>x. f x + g x) absolutely_integrable_on s"
+proof -
+  let ?P = "\<And>f g::'n \<Rightarrow> 'm. f absolutely_integrable_on UNIV \<Longrightarrow>
+    g absolutely_integrable_on UNIV \<Longrightarrow> (\<lambda>x. f x + g x) absolutely_integrable_on UNIV"
+  {
+    presume as: "PROP ?P"
+    note a = absolutely_integrable_restrict_univ[symmetric]
+    have *: "\<And>x. (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0) =
+      (if x \<in> s then f x + g x else 0)" by auto
+    show ?thesis
+      apply (subst a)
+      using as[OF assms[unfolded a[of f] a[of g]]]
+      apply (simp only: *)
+      done
+  }
+  fix f g :: "'n \<Rightarrow> 'm"
+  assume assms: "f absolutely_integrable_on UNIV" "g absolutely_integrable_on UNIV"
+  note absolutely_integrable_bounded_variation
+  from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
+  show "(\<lambda>x. f x + g x) absolutely_integrable_on UNIV"
+    apply (rule bounded_variation_absolutely_integrable[of _ "B1+B2"])
+    apply (rule integrable_add)
+    prefer 3
+    apply safe
+  proof goal_cases
+    case prems: (1 d)
+    have "\<And>k. k \<in> d \<Longrightarrow> f integrable_on k \<and> g integrable_on k"
+      apply (drule division_ofD(4)[OF prems])
+      apply safe
+      apply (rule_tac[!] integrable_on_subcbox[of _ UNIV])
+      using assms
+      apply auto
+      done
+    then have "(\<Sum>k\<in>d. norm (integral k (\<lambda>x. f x + g x))) \<le>
+      (\<Sum>k\<in>d. norm (integral k f)) + (\<Sum>k\<in>d. norm (integral k g))"
+      apply -
+      unfolding setsum.distrib [symmetric]
+      apply (rule setsum_mono)
+      apply (subst integral_add)
+      prefer 3
+      apply (rule norm_triangle_ineq)
+      apply auto
+      done
+    also have "\<dots> \<le> B1 + B2"
+      using B(1)[OF prems] B(2)[OF prems] by auto
+    finally show ?case .
+  qed (insert assms, auto)
+qed
+
+lemma absolutely_integrable_sub[intro]:
+  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "f absolutely_integrable_on s"
+    and "g absolutely_integrable_on s"
+  shows "(\<lambda>x. f x - g x) absolutely_integrable_on s"
+  using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
+  by (simp add: algebra_simps)
+
+lemma absolutely_integrable_linear:
+  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
+    and h :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
+  assumes "f absolutely_integrable_on s"
+    and "bounded_linear h"
+  shows "(h \<circ> f) absolutely_integrable_on s"
+proof -
+  {
+    presume as: "\<And>f::'m \<Rightarrow> 'n. \<And>h::'n \<Rightarrow> 'p. f absolutely_integrable_on UNIV \<Longrightarrow>
+      bounded_linear h \<Longrightarrow> (h \<circ> f) absolutely_integrable_on UNIV"
+    note a = absolutely_integrable_restrict_univ[symmetric]
+    show ?thesis
+      apply (subst a)
+      using as[OF assms[unfolded a[of f] a[of g]]]
+      apply (simp only: o_def if_distrib linear_simps[OF assms(2)])
+      done
+  }
+  fix f :: "'m \<Rightarrow> 'n"
+  fix h :: "'n \<Rightarrow> 'p"
+  assume assms: "f absolutely_integrable_on UNIV" "bounded_linear h"
+  from absolutely_integrable_bounded_variation[OF assms(1)] guess B . note B=this
+  from bounded_linear.pos_bounded[OF assms(2)] guess b .. note b=conjunctD2[OF this]
+  show "(h \<circ> f) absolutely_integrable_on UNIV"
+    apply (rule bounded_variation_absolutely_integrable[of _ "B * b"])
+    apply (rule integrable_linear[OF _ assms(2)])
+    apply safe
+  proof goal_cases
+    case prems: (2 d)
+    have "(\<Sum>k\<in>d. norm (integral k (h \<circ> f))) \<le> setsum (\<lambda>k. norm(integral k f)) d * b"
+      unfolding setsum_left_distrib
+      apply (rule setsum_mono)
+    proof goal_cases
+      case (1 k)
+      from division_ofD(4)[OF prems this]
+      guess u v by (elim exE) note uv=this
+      have *: "f integrable_on k"
+        unfolding uv
+        apply (rule integrable_on_subcbox[of _ UNIV])
+        using assms
+        apply auto
+        done
+      note this[unfolded has_integral_integral]
+      note has_integral_linear[OF this assms(2)] integrable_linear[OF * assms(2)]
+      note * = has_integral_unique[OF this(2)[unfolded has_integral_integral] this(1)]
+      show ?case
+        unfolding * using b by auto
+    qed
+    also have "\<dots> \<le> B * b"
+      apply (rule mult_right_mono)
+      using B prems b
+      apply auto
+      done
+    finally show ?case .
+  qed (insert assms, auto)
+qed
+
+lemma absolutely_integrable_setsum:
+  fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "finite t"
+    and "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s"
+  shows "(\<lambda>x. setsum (\<lambda>a. f a x) t) absolutely_integrable_on s"
+  using assms(1,2)
+  by induct auto
+
+lemma absolutely_integrable_vector_abs:
+  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
+    and T :: "'c::euclidean_space \<Rightarrow> 'b"
+  assumes f: "f absolutely_integrable_on s"
+  shows "(\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>T i\<bar> *\<^sub>R i)) absolutely_integrable_on s"
+  (is "?Tf absolutely_integrable_on s")
+proof -
+  have if_distrib: "\<And>P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)"
+    by simp
+  have *: "\<And>x. ?Tf x = (\<Sum>i\<in>Basis.
+    ((\<lambda>y. (\<Sum>j\<in>Basis. (if j = i then y else 0) *\<^sub>R j)) o
+     (\<lambda>x. (norm (\<Sum>j\<in>Basis. (if j = i then f x\<bullet>T i else 0) *\<^sub>R j)))) x)"
+    by (simp add: comp_def if_distrib setsum.If_cases)
+  show ?thesis
+    unfolding *
+    apply (rule absolutely_integrable_setsum[OF finite_Basis])
+    apply (rule absolutely_integrable_linear)
+    apply (rule absolutely_integrable_norm)
+    apply (rule absolutely_integrable_linear[OF f, unfolded o_def])
+    apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI)
+    done
+qed
+
+lemma absolutely_integrable_max:
+  fixes f g :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
+  assumes "f absolutely_integrable_on s"
+    and "g absolutely_integrable_on s"
+  shows "(\<lambda>x. (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
+proof -
+  have *:"\<And>x. (1 / 2) *\<^sub>R (((\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i)::'n) + (f x + g x)) =
+      (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
+    unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
+  note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms]
+  note absolutely_integrable_vector_abs[OF this(1), where T="\<lambda>x. x"] this(2)
+  note absolutely_integrable_add[OF this]
+  note absolutely_integrable_cmul[OF this, of "1/2"]
+  then show ?thesis unfolding * .
+qed
+
+lemma absolutely_integrable_min:
+  fixes f g::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
+  assumes "f absolutely_integrable_on s"
+    and "g absolutely_integrable_on s"
+  shows "(\<lambda>x. (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
+proof -
+  have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - (\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i::'n)) =
+      (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
+    unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
+
+  note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms]
+  note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\<lambda>x. x"]
+  note absolutely_integrable_sub[OF this]
+  note absolutely_integrable_cmul[OF this,of "1/2"]
+  then show ?thesis unfolding * .
+qed
+
+lemma absolutely_integrable_abs_eq:
+  fixes f::"'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and>
+    (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on s"
+  (is "?l = ?r")
+proof
+  assume ?l
+  then show ?r
+    apply -
+    apply rule
+    defer
+    apply (drule absolutely_integrable_vector_abs)
+    apply auto
+    done
+next
+  assume ?r
+  {
+    presume lem: "\<And>f::'n \<Rightarrow> 'm. f integrable_on UNIV \<Longrightarrow>
+      (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV \<Longrightarrow>
+        f absolutely_integrable_on UNIV"
+    have *: "\<And>x. (\<Sum>i\<in>Basis. \<bar>(if x \<in> s then f x else 0) \<bullet> i\<bar> *\<^sub>R i) =
+      (if x \<in> s then (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) else (0::'m))"
+      unfolding euclidean_eq_iff[where 'a='m]
+      by auto
+    show ?l
+      apply (subst absolutely_integrable_restrict_univ[symmetric])
+      apply (rule lem)
+      unfolding integrable_restrict_univ *
+      using \<open>?r\<close>
+      apply auto
+      done
+  }
+  fix f :: "'n \<Rightarrow> 'm"
+  assume assms: "f integrable_on UNIV" "(\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV"
+  let ?B = "\<Sum>i\<in>Basis. integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
+  show "f absolutely_integrable_on UNIV"
+    apply (rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"])
+    apply safe
+  proof goal_cases
+    case d: (1 d)
+    note d'=division_ofD[OF d]
+    have "(\<Sum>k\<in>d. norm (integral k f)) \<le>
+      (\<Sum>k\<in>d. setsum (op \<bullet> (integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis)"
+      apply (rule setsum_mono)
+      apply (rule order_trans[OF norm_le_l1])
+      apply (rule setsum_mono)
+      unfolding lessThan_iff
+    proof -
+      fix k
+      fix i :: 'm
+      assume "k \<in> d" and i: "i \<in> Basis"
+      from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this
+      show "\<bar>integral k f \<bullet> i\<bar> \<le> integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
+        apply (rule abs_leI)
+        unfolding inner_minus_left[symmetric]
+        defer
+        apply (subst integral_neg[symmetric])
+        apply (rule_tac[1-2] integral_component_le[OF i])
+        using integrable_on_subcbox[OF assms(1),of a b]
+          integrable_on_subcbox[OF assms(2),of a b] i  integrable_neg
+        unfolding ab
+        apply auto
+        done
+    qed
+    also have "\<dots> \<le> setsum (op \<bullet> (integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis"
+      apply (subst setsum.commute)
+      apply (rule setsum_mono)
+    proof goal_cases
+      case (1 j)
+      have *: "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) integrable_on \<Union>d"
+        using integrable_on_subdivision[OF d assms(2)] by auto
+      have "(\<Sum>i\<in>d. integral i (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j) =
+        integral (\<Union>d) (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
+        unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] ..
+      also have "\<dots> \<le> integral UNIV (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
+        apply (rule integral_subset_component_le)
+        using assms * \<open>j \<in> Basis\<close>
+        apply auto
+        done
+      finally show ?case .
+    qed
+    finally show ?case .
+  qed
+qed
+
+lemma nonnegative_absolutely_integrable:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "\<forall>x\<in>s. \<forall>i\<in>Basis. 0 \<le> f x \<bullet> i"
+    and "f integrable_on s"
+  shows "f absolutely_integrable_on s"
+  unfolding absolutely_integrable_abs_eq
+  apply rule
+  apply (rule assms)thm integrable_eq
+  apply (rule integrable_eq[of _ f])
+  using assms
+  apply (auto simp: euclidean_eq_iff[where 'a='m])
+  done
+
+lemma absolutely_integrable_integrable_bound:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "\<forall>x\<in>s. norm (f x) \<le> g x"
+    and "f integrable_on s"
+    and "g integrable_on s"
+  shows "f absolutely_integrable_on s"
+proof -
+  {
+    presume *: "\<And>f::'n \<Rightarrow> 'm. \<And>g. \<forall>x. norm (f x) \<le> g x \<Longrightarrow> f integrable_on UNIV \<Longrightarrow>
+      g integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
+    show ?thesis
+      apply (subst absolutely_integrable_restrict_univ[symmetric])
+      apply (rule *[of _ "\<lambda>x. if x\<in>s then g x else 0"])
+      using assms
+      unfolding integrable_restrict_univ
+      apply auto
+      done
+  }
+  fix g
+  fix f :: "'n \<Rightarrow> 'm"
+  assume assms: "\<forall>x. norm (f x) \<le> g x" "f integrable_on UNIV" "g integrable_on UNIV"
+  show "f absolutely_integrable_on UNIV"
+    apply (rule bounded_variation_absolutely_integrable[OF assms(2),where B="integral UNIV g"])
+    apply safe
+  proof goal_cases
+    case d: (1 d)
+    note d'=division_ofD[OF d]
+    have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>k\<in>d. integral k g)"
+      apply (rule setsum_mono)
+      apply (rule integral_norm_bound_integral)
+      apply (drule_tac[!] d'(4))
+      apply safe
+      apply (rule_tac[1-2] integrable_on_subcbox)
+      using assms
+      apply auto
+      done
+    also have "\<dots> = integral (\<Union>d) g"
+      apply (rule integral_combine_division_bottomup[symmetric])
+      apply (rule d)
+      apply safe
+      apply (drule d'(4))
+      apply safe
+      apply (rule integrable_on_subcbox[OF assms(3)])
+      apply auto
+      done
+    also have "\<dots> \<le> integral UNIV g"
+      apply (rule integral_subset_le)
+      defer
+      apply (rule integrable_on_subdivision[OF d,of _ UNIV])
+      prefer 4
+      apply rule
+      apply (rule_tac y="norm (f x)" in order_trans)
+      using assms
+      apply auto
+      done
+    finally show ?case .
+  qed
+qed
+
+lemma absolutely_integrable_absolutely_integrable_bound:
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+    and g :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
+  assumes "\<forall>x\<in>s. norm (f x) \<le> norm (g x)"
+    and "f integrable_on s"
+    and "g absolutely_integrable_on s"
+  shows "f absolutely_integrable_on s"
+  apply (rule absolutely_integrable_integrable_bound[of s f "\<lambda>x. norm (g x)"])
+  using assms
+  apply auto
+  done
+
+lemma absolutely_integrable_inf_real:
+  assumes "finite k"
+    and "k \<noteq> {}"
+    and "\<forall>i\<in>k. (\<lambda>x. (fs x i)::real) absolutely_integrable_on s"
+  shows "(\<lambda>x. (Inf ((fs x) ` k))) absolutely_integrable_on s"
+  using assms
+proof induct
+  case (insert a k)
+  let ?P = "(\<lambda>x.
+    if fs x ` k = {} then fs x a
+    else min (fs x a) (Inf (fs x ` k))) absolutely_integrable_on s"
+  show ?case
+    unfolding image_insert
+    apply (subst Inf_insert_finite)
+    apply (rule finite_imageI[OF insert(1)])
+  proof (cases "k = {}")
+    case True
+    then show ?P
+      apply (subst if_P)
+      defer
+      apply (rule insert(5)[rule_format])
+      apply auto
+      done
+  next
+    case False
+    then show ?P
+      apply (subst if_not_P)
+      defer
+      apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified])
+      defer
+      apply(rule insert(3)[OF False])
+      using insert(5)
+      apply auto
+      done
+  qed
+next
+  case empty
+  then show ?case by auto
+qed
+
+lemma absolutely_integrable_sup_real:
+  assumes "finite k"
+    and "k \<noteq> {}"
+    and "\<forall>i\<in>k. (\<lambda>x. (fs x i)::real) absolutely_integrable_on s"
+  shows "(\<lambda>x. (Sup ((fs x) ` k))) absolutely_integrable_on s"
+  using assms
+proof induct
+  case (insert a k)
+  let ?P = "(\<lambda>x.
+    if fs x ` k = {} then fs x a
+    else max (fs x a) (Sup (fs x ` k))) absolutely_integrable_on s"
+  show ?case
+    unfolding image_insert
+    apply (subst Sup_insert_finite)
+    apply (rule finite_imageI[OF insert(1)])
+  proof (cases "k = {}")
+    case True
+    then show ?P
+      apply (subst if_P)
+      defer
+      apply (rule insert(5)[rule_format])
+      apply auto
+      done
+  next
+    case False
+    then show ?P
+      apply (subst if_not_P)
+      defer
+      apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified])
+      defer
+      apply (rule insert(3)[OF False])
+      using insert(5)
+      apply auto
+      done
+  qed
+qed auto
+
+
+subsection \<open>differentiation under the integral sign\<close>
+
+lemma tube_lemma:
+  assumes "compact K"
+  assumes "open W"
+  assumes "{x0} \<times> K \<subseteq> W"
+  shows "\<exists>X0. x0 \<in> X0 \<and> open X0 \<and> X0 \<times> K \<subseteq> W"
+proof -
+  {
+    fix y assume "y \<in> K"
+    then have "(x0, y) \<in> W" using assms by auto
+    with \<open>open W\<close>
+    have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W"
+      by (rule open_prod_elim) blast
+  }
+  then obtain X0 Y where
+    *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
+    by metis
+  from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
+  with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
+    by (rule compactE)
+  then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
+    by (force intro!: choice)
+  with * CC show ?thesis
+    by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"]) (* SLOW *)
+qed
+
+lemma continuous_on_prod_compactE:
+  fixes fx::"'a::topological_space \<times> 'b::topological_space \<Rightarrow> 'c::metric_space"
+    and e::real
+  assumes cont_fx: "continuous_on (U \<times> C) fx"
+  assumes "compact C"
+  assumes [intro]: "x0 \<in> U"
+  notes [continuous_intros] = continuous_on_compose2[OF cont_fx]
+  assumes "e > 0"
+  obtains X0 where "x0 \<in> X0" "open X0"
+    "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
+proof -
+  define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))"
+  define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
+  have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
+    by (auto simp: vimage_def W0_def)
+  have "open {..<e}" by simp
+  have "continuous_on (U \<times> C) psi"
+    by (auto intro!: continuous_intros simp: psi_def split_beta')
+  from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
+  obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
+    unfolding W0_eq by blast
+  have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
+    unfolding W
+    by (auto simp: W0_def psi_def \<open>0 < e\<close>)
+  then have "{x0} \<times> C \<subseteq> W" by blast
+  from tube_lemma[OF \<open>compact C\<close> \<open>open W\<close> this]
+  obtain X0 where X0: "x0 \<in> X0" "open X0" "X0 \<times> C \<subseteq> W"
+    by blast
+
+  have "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
+  proof safe
+    fix x assume x: "x \<in> X0" "x \<in> U"
+    fix t assume t: "t \<in> C"
+    have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)"
+      by (auto simp: psi_def)
+    also
+    {
+      have "(x, t) \<in> X0 \<times> C"
+        using t x
+        by auto
+      also note \<open>\<dots> \<subseteq> W\<close>
+      finally have "(x, t) \<in> W" .
+      with t x have "(x, t) \<in> W \<inter> U \<times> C"
+        by blast
+      also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close>
+      finally  have "psi (x, t) < e"
+        by (auto simp: W0_def)
+    }
+    finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp
+  qed
+  from X0(1,2) this show ?thesis ..
+qed
+
+lemma integral_continuous_on_param:
+  fixes f::"'a::topological_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
+  assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). f x t)"
+  shows "continuous_on U (\<lambda>x. integral (cbox a b) (f x))"
+proof cases
+  assume "content (cbox a b) \<noteq> 0"
+  then have ne: "cbox a b \<noteq> {}" by auto
+
+  note [continuous_intros] =
+    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
+      unfolded split_beta fst_conv snd_conv]
+
+  show ?thesis
+    unfolding continuous_on_def
+  proof (safe intro!: tendstoI)
+    fix e'::real and x
+    assume "e' > 0"
+    define e where "e = e' / (content (cbox a b) + 1)"
+    have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
+    assume "x \<in> U"
+    from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x \<in> U\<close> \<open>0 < e\<close>]
+    obtain X0 where X0: "x \<in> X0" "open X0"
+      and fx_bound: "\<And>y t. y \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (f y t - f x t) \<le> e"
+      unfolding split_beta fst_conv snd_conv dist_norm
+      by metis
+    have "\<forall>\<^sub>F y in at x within U. y \<in> X0 \<inter> U"
+      using X0(1) X0(2) eventually_at_topological by auto
+    then show "\<forall>\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'"
+    proof eventually_elim
+      case (elim y)
+      have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) =
+        norm (integral (cbox a b) (\<lambda>t. f y t - f x t))"
+        using elim \<open>x \<in> U\<close>
+        unfolding dist_norm
+        by (subst integral_diff)
+           (auto intro!: integrable_continuous continuous_intros)
+      also have "\<dots> \<le> e * content (cbox a b)"
+        using elim \<open>x \<in> U\<close>
+        by (intro integrable_bound)
+           (auto intro!: fx_bound \<open>x \<in> U \<close> less_imp_le[OF \<open>0 < e\<close>]
+              integrable_continuous continuous_intros)
+      also have "\<dots> < e'"
+        using \<open>0 < e'\<close> \<open>e > 0\<close>
+        by (auto simp: e_def divide_simps)
+      finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" .
+    qed
+  qed
+qed (auto intro!: continuous_on_const)
+
+lemma eventually_closed_segment:
+  fixes x0::"'a::real_normed_vector"
+  assumes "open X0" "x0 \<in> X0"
+  shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
+proof -
+  from openE[OF assms]
+  obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
+  then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
+    by (auto simp: dist_commute eventually_at)
+  then show ?thesis
+  proof eventually_elim
+    case (elim x)
+    have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
+    from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
+    have "closed_segment x0 x \<subseteq> ball x0 e" .
+    also note \<open>\<dots> \<subseteq> X0\<close>
+    finally show ?case .
+  qed
+qed
+
+lemma leibniz_rule:
+  fixes f::"'a::banach \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
+  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
+    ((\<lambda>x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)"
+  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> f x integrable_on cbox a b"
+  assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
+  assumes [intro]: "x0 \<in> U"
+  assumes "convex U"
+  shows
+    "((\<lambda>x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
+    (is "(?F has_derivative ?dF) _")
+proof cases
+  assume "content (cbox a b) \<noteq> 0"
+  then have ne: "cbox a b \<noteq> {}" by auto
+  note [continuous_intros] =
+    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
+      unfolded split_beta fst_conv snd_conv]
+  show ?thesis
+  proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist)
+    have cont_f1: "\<And>t. t \<in> cbox a b \<Longrightarrow> continuous_on U (\<lambda>x. f x t)"
+      by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx)
+    note [continuous_intros] = continuous_on_compose2[OF cont_f1]
+    fix e'::real
+    assume "e' > 0"
+    define e where "e = e' / (content (cbox a b) + 1)"
+    have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
+    from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x0 \<in> U\<close> \<open>e > 0\<close>]
+    obtain X0 where X0: "x0 \<in> X0" "open X0"
+      and fx_bound: "\<And>x t. x \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (fx x t - fx x0 t) \<le> e"
+      unfolding split_beta fst_conv snd_conv
+      by (metis dist_norm)
+
+    note eventually_closed_segment[OF \<open>open X0\<close> \<open>x0 \<in> X0\<close>, of U]
+    moreover
+    have "\<forall>\<^sub>F x in at x0 within U. x \<in> X0"
+      using \<open>open X0\<close> \<open>x0 \<in> X0\<close> eventually_at_topological by blast
+    moreover have "\<forall>\<^sub>F x in at x0 within U. x \<noteq> x0"
+      by (auto simp: eventually_at_filter)
+    moreover have "\<forall>\<^sub>F x in at x0 within U. x \<in> U"
+      by (auto simp: eventually_at_filter)
+    ultimately
+    show "\<forall>\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'"
+    proof eventually_elim
+      case (elim x)
+      from elim have "0 < norm (x - x0)" by simp
+      have "closed_segment x0 x \<subseteq> U"
+        by (rule \<open>convex U\<close>[unfolded convex_contains_segment, rule_format, OF \<open>x0 \<in> U\<close> \<open>x \<in> U\<close>])
+      from elim have [intro]: "x \<in> U" by auto
+
+      have "?F x - ?F x0 - ?dF (x - x0) =
+        integral (cbox a b) (\<lambda>y. f x y - f x0 y - fx x0 y (x - x0))"
+        (is "_ = ?id")
+        using \<open>x \<noteq> x0\<close>
+        by (subst blinfun_apply_integral integral_diff,
+            auto intro!: integrable_diff integrable_f2 continuous_intros
+              intro: integrable_continuous)+
+      also
+      {
+        fix t assume t: "t \<in> (cbox a b)"
+        have seg: "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R (x - x0) \<in> X0 \<inter> U"
+          using \<open>closed_segment x0 x \<subseteq> U\<close>
+            \<open>closed_segment x0 x \<subseteq> X0\<close>
+          by (force simp: closed_segment_def algebra_simps)
+        from t have deriv:
+          "((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)"
+          if "y \<in> X0 \<inter> U" for y
+          unfolding has_vector_derivative_def[symmetric]
+          using that \<open>x \<in> X0\<close>
+          by (intro has_derivative_within_subset[OF fx]) auto
+        have "\<forall>x \<in> X0 \<inter> U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \<le> e"
+          using fx_bound t
+          by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric])
+        from differentiable_bound_linearization[OF seg deriv this] X0
+        have "norm (f x t - f x0 t - fx x0 t (x - x0)) \<le> e * norm (x - x0)"
+          by (auto simp add: ac_simps)
+      }
+      then have "norm ?id \<le> integral (cbox a b) (\<lambda>_. e * norm (x - x0))"
+        by (intro integral_norm_bound_integral)
+          (auto intro!: continuous_intros integrable_diff integrable_f2
+            intro: integrable_continuous)
+      also have "\<dots> = content (cbox a b) * e * norm (x - x0)"
+        by simp
+      also have "\<dots> < e' * norm (x - x0)"
+        using \<open>e' > 0\<close> content_pos_le[of a b]
+        by (intro mult_strict_right_mono[OF _ \<open>0 < norm (x - x0)\<close>])
+          (auto simp: divide_simps e_def)
+      finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" .
+      then show ?case
+        by (auto simp: divide_simps)
+    qed
+  qed (rule blinfun.bounded_linear_right)
+qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps)
+
+lemma
+  has_vector_derivative_eq_has_derivative_blinfun:
+  "(f has_vector_derivative f') (at x within U) \<longleftrightarrow>
+    (f has_derivative blinfun_scaleR_left f') (at x within U)"
+  by (simp add: has_vector_derivative_def)
+
+lemma leibniz_rule_vector_derivative:
+  fixes f::"real \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
+  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
+      ((\<lambda>x. f x t) has_vector_derivative (fx x t)) (at x within U)"
+  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
+  assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). fx x t)"
+  assumes U: "x0 \<in> U" "convex U"
+  shows "((\<lambda>x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0))
+      (at x0 within U)"
+proof -
+  note [continuous_intros] =
+    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
+      unfolded split_beta fst_conv snd_conv]
+  have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) =
+    integral (cbox a b) (\<lambda>t. blinfun_scaleR_left (fx x0 t))"
+    by (subst integral_linear[symmetric])
+       (auto simp: has_vector_derivative_def o_def
+         intro!: integrable_continuous U continuous_intros bounded_linear_intros)
+  show ?thesis
+    unfolding has_vector_derivative_eq_has_derivative_blinfun
+    apply (rule has_derivative_eq_rhs)
+    apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_scaleR_left (fx x t)"])
+    using fx cont_fx
+    apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros)
+    done
+qed
+
+lemma
+  has_field_derivative_eq_has_derivative_blinfun:
+  "(f has_field_derivative f') (at x within U) \<longleftrightarrow> (f has_derivative blinfun_mult_right f') (at x within U)"
+  by (simp add: has_field_derivative_def)
+
+lemma leibniz_rule_field_derivative:
+  fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
+  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)"
+  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
+  assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
+  assumes U: "x0 \<in> U" "convex U"
+  shows "((\<lambda>x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
+proof -
+  note [continuous_intros] =
+    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
+      unfolded split_beta fst_conv snd_conv]
+  have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) =
+    integral (cbox a b) (\<lambda>t. blinfun_mult_right (fx x0 t))"
+    by (subst integral_linear[symmetric])
+      (auto simp: has_vector_derivative_def o_def
+        intro!: integrable_continuous U continuous_intros bounded_linear_intros)
+  show ?thesis
+    unfolding has_field_derivative_eq_has_derivative_blinfun
+    apply (rule has_derivative_eq_rhs)
+    apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_mult_right (fx x t)"])
+    using fx cont_fx
+    apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros)
+    done
+qed
+
+
+subsection \<open>Exchange uniform limit and integral\<close>
+
+lemma
+  uniform_limit_integral:
+  fixes f::"'a \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
+  assumes u: "uniform_limit (cbox a b) f g F"
+  assumes c: "\<And>n. continuous_on (cbox a b) (f n)"
+  assumes [simp]: "F \<noteq> bot"
+  obtains I J where
+    "\<And>n. (f n has_integral I n) (cbox a b)"
+    "(g has_integral J) (cbox a b)"
+    "(I \<longlongrightarrow> J) F"
+proof -
+  have fi[simp]: "f n integrable_on (cbox a b)" for n
+    by (auto intro!: integrable_continuous assms)
+  then obtain I where I: "\<And>n. (f n has_integral I n) (cbox a b)"
+    by atomize_elim (auto simp: integrable_on_def intro!: choice)
+
+  moreover
+
+  have gi[simp]: "g integrable_on (cbox a b)"
+    by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c)
+  then obtain J where J: "(g has_integral J) (cbox a b)"
+    by blast
+
+  moreover
+
+  have "(I \<longlongrightarrow> J) F"
+  proof cases
+    assume "content (cbox a b) = 0"
+    hence "I = (\<lambda>_. 0)" "J = 0"
+      by (auto intro!: has_integral_unique I J)
+    thus ?thesis by simp
+  next
+    assume content_nonzero: "content (cbox a b) \<noteq> 0"
+    show ?thesis
+    proof (rule tendstoI)
+      fix e::real
+      assume "e > 0"
+      define e' where "e' = e / 2"
+      with \<open>e > 0\<close> have "e' > 0" by simp
+      then have "\<forall>\<^sub>F n in F. \<forall>x\<in>cbox a b. norm (f n x - g x) < e' / content (cbox a b)"
+        using u content_nonzero content_pos_le[of a b]
+        by (auto simp: uniform_limit_iff dist_norm)
+      then show "\<forall>\<^sub>F n in F. dist (I n) J < e"
+      proof eventually_elim
+        case (elim n)
+        have "I n = integral (cbox a b) (f n)"
+            "J = integral (cbox a b) g"
+          using I[of n] J by (simp_all add: integral_unique)
+        then have "dist (I n) J = norm (integral (cbox a b) (\<lambda>x. f n x - g x))"
+          by (simp add: integral_diff dist_norm)
+        also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. (e' / content (cbox a b)))"
+          using elim
+          by (intro integral_norm_bound_integral)
+            (auto intro!: integrable_diff absolutely_integrable_onI)
+        also have "\<dots> < e"
+          using \<open>0 < e\<close>
+          by (simp add: e'_def)
+        finally show ?case .
+      qed
+    qed
+  qed
+  ultimately show ?thesis ..
+qed
+
+
+subsection \<open>Dominated convergence\<close>
+
+context
+begin
+
+private lemma dominated_convergence_real:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
+  assumes "\<And>k. (f k) integrable_on s" "h integrable_on s"
+    and "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
+    and "\<forall>x \<in> s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+proof
+  have bdd_below[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_below {f n x |n. P n}"
+  proof (safe intro!: bdd_belowI)
+    fix n x show "x \<in> s \<Longrightarrow> - h x \<le> f n x"
+      using assms(3)[rule_format, of x n] by simp
+  qed
+  have bdd_above[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_above {f n x |n. P n}"
+  proof (safe intro!: bdd_aboveI)
+    fix n x show "x \<in> s \<Longrightarrow> f n x \<le> h x"
+      using assms(3)[rule_format, of x n] by simp
+  qed
+
+  have "\<And>m. (\<lambda>x. Inf {f j x |j. m \<le> j}) integrable_on s \<and>
+    ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) \<longlongrightarrow>
+    integral s (\<lambda>x. Inf {f j x |j. m \<le> j}))sequentially"
+  proof (rule monotone_convergence_decreasing, safe)
+    fix m :: nat
+    show "bounded {integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}}) |k. True}"
+      unfolding bounded_iff
+      apply (rule_tac x="integral s h" in exI)
+    proof safe
+      fix k :: nat
+      show "norm (integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) \<le> integral s h"
+        apply (rule integral_norm_bound_integral)
+        unfolding simple_image
+        apply (rule absolutely_integrable_onD)
+        apply (rule absolutely_integrable_inf_real)
+        prefer 5
+        unfolding real_norm_def
+        apply rule
+        apply (rule cInf_abs_ge)
+        prefer 5
+        apply rule
+        apply (rule_tac g = h in absolutely_integrable_integrable_bound)
+        using assms
+        unfolding real_norm_def
+        apply auto
+        done
+    qed
+    fix k :: nat
+    show "(\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}}) integrable_on s"
+      unfolding simple_image
+      apply (rule absolutely_integrable_onD)
+      apply (rule absolutely_integrable_inf_real)
+      prefer 3
+      using absolutely_integrable_integrable_bound[OF assms(3,1,2)]
+      apply auto
+      done
+    fix x
+    assume x: "x \<in> s"
+    show "Inf {f j x |j. j \<in> {m..m + Suc k}} \<le> Inf {f j x |j. j \<in> {m..m + k}}"
+      by (rule cInf_superset_mono) auto
+    let ?S = "{f j x| j. m \<le> j}"
+    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) \<longlongrightarrow> Inf ?S) sequentially"
+    proof (rule LIMSEQ_I, goal_cases)
+      case r: (1 r)
+
+      have "\<exists>y\<in>?S. y < Inf ?S + r"
+        by (subst cInf_less_iff[symmetric]) (auto simp: \<open>x\<in>s\<close> r)
+      then obtain N where N: "f N x < Inf ?S + r" "m \<le> N"
+        by blast
+
+      show ?case
+        apply (rule_tac x=N in exI)
+        apply safe
+      proof goal_cases
+        case prems: (1 n)
+        have *: "\<And>y ix. y < Inf ?S + r \<longrightarrow> Inf ?S \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> \<bar>ix - Inf ?S\<bar> < r"
+          by arith
+        show ?case
+          unfolding real_norm_def
+            apply (rule *[rule_format, OF N(1)])
+            apply (rule cInf_superset_mono, auto simp: \<open>x\<in>s\<close>) []
+            apply (rule cInf_lower)
+            using N prems
+            apply auto []
+            apply simp
+            done
+      qed
+    qed
+  qed
+  note dec1 = conjunctD2[OF this]
+
+  have "\<And>m. (\<lambda>x. Sup {f j x |j. m \<le> j}) integrable_on s \<and>
+    ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}})) \<longlongrightarrow>
+    integral s (\<lambda>x. Sup {f j x |j. m \<le> j})) sequentially"
+  proof (rule monotone_convergence_increasing,safe)
+    fix m :: nat
+    show "bounded {integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}}) |k. True}"
+      unfolding bounded_iff
+      apply (rule_tac x="integral s h" in exI)
+    proof safe
+      fix k :: nat
+      show "norm (integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}})) \<le> integral s h"
+        apply (rule integral_norm_bound_integral) unfolding simple_image
+        apply (rule absolutely_integrable_onD)
+        apply(rule absolutely_integrable_sup_real)
+        prefer 5 unfolding real_norm_def
+        apply rule
+        apply (rule cSup_abs_le)
+        using assms
+        apply (force simp add: )
+        prefer 4
+        apply rule
+        apply (rule_tac g=h in absolutely_integrable_integrable_bound)
+        using assms
+        unfolding real_norm_def
+        apply auto
+        done
+    qed
+    fix k :: nat
+    show "(\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}}) integrable_on s"
+      unfolding simple_image
+      apply (rule absolutely_integrable_onD)
+      apply (rule absolutely_integrable_sup_real)
+      prefer 3
+      using absolutely_integrable_integrable_bound[OF assms(3,1,2)]
+      apply auto
+      done
+    fix x
+    assume x: "x\<in>s"
+    show "Sup {f j x |j. j \<in> {m..m + Suc k}} \<ge> Sup {f j x |j. j \<in> {m..m + k}}"
+      by (rule cSup_subset_mono) auto
+    let ?S = "{f j x| j. m \<le> j}"
+    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) \<longlongrightarrow> Sup ?S) sequentially"
+    proof (rule LIMSEQ_I, goal_cases)
+      case r: (1 r)
+      have "\<exists>y\<in>?S. Sup ?S - r < y"
+        by (subst less_cSup_iff[symmetric]) (auto simp: r \<open>x\<in>s\<close>)
+      then obtain N where N: "Sup ?S - r < f N x" "m \<le> N"
+        by blast
+
+      show ?case
+        apply (rule_tac x=N in exI)
+        apply safe
+      proof goal_cases
+        case prems: (1 n)
+        have *: "\<And>y ix. Sup ?S - r < y \<longrightarrow> ix \<le> Sup ?S \<longrightarrow> y \<le> ix \<longrightarrow> \<bar>ix - Sup ?S\<bar> < r"
+          by arith
+        show ?case
+          apply simp
+          apply (rule *[rule_format, OF N(1)])
+          apply (rule cSup_subset_mono, auto simp: \<open>x\<in>s\<close>) []
+          apply (subst cSup_upper)
+          using N prems
+          apply auto
+          done
+      qed
+    qed
+  qed
+  note inc1 = conjunctD2[OF this]
+
+  have "g integrable_on s \<and>
+    ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. k \<le> j})) \<longlongrightarrow> integral s g) sequentially"
+    apply (rule monotone_convergence_increasing,safe)
+    apply fact
+  proof -
+    show "bounded {integral s (\<lambda>x. Inf {f j x |j. k \<le> j}) |k. True}"
+      unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
+    proof safe
+      fix k :: nat
+      show "norm (integral s (\<lambda>x. Inf {f j x |j. k \<le> j})) \<le> integral s h"
+        apply (rule integral_norm_bound_integral)
+        apply fact+
+        unfolding real_norm_def
+        apply rule
+        apply (rule cInf_abs_ge)
+        using assms(3)
+        apply auto
+        done
+    qed
+    fix k :: nat and x
+    assume x: "x \<in> s"
+
+    have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
+    show "Inf {f j x |j. k \<le> j} \<le> Inf {f j x |j. Suc k \<le> j}"
+      by (intro cInf_superset_mono) (auto simp: \<open>x\<in>s\<close>)
+
+    show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) \<longlonglongrightarrow> g x"
+    proof (rule LIMSEQ_I, goal_cases)
+      case r: (1 r)
+      then have "0<r/2"
+        by auto
+      from assms(4)[THEN bspec, THEN LIMSEQ_D, OF x this] guess N .. note N = this
+      show ?case
+        apply (rule_tac x=N in exI)
+        apply safe
+        unfolding real_norm_def
+        apply (rule le_less_trans[of _ "r/2"])
+        apply (rule cInf_asclose)
+        apply safe
+        defer
+        apply (rule less_imp_le)
+        using N r
+        apply auto
+        done
+    qed
+  qed
+  note inc2 = conjunctD2[OF this]
+
+  have "g integrable_on s \<and>
+    ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. k \<le> j})) \<longlongrightarrow> integral s g) sequentially"
+    apply (rule monotone_convergence_decreasing,safe)
+    apply fact
+  proof -
+    show "bounded {integral s (\<lambda>x. Sup {f j x |j. k \<le> j}) |k. True}"
+      unfolding bounded_iff
+      apply (rule_tac x="integral s h" in exI)
+    proof safe
+      fix k :: nat
+      show "norm (integral s (\<lambda>x. Sup {f j x |j. k \<le> j})) \<le> integral s h"
+        apply (rule integral_norm_bound_integral)
+        apply fact+
+        unfolding real_norm_def
+        apply rule
+        apply (rule cSup_abs_le)
+        using assms(3)
+        apply auto
+        done
+    qed
+    fix k :: nat
+    fix x
+    assume x: "x \<in> s"
+
+    show "Sup {f j x |j. k \<le> j} \<ge> Sup {f j x |j. Suc k \<le> j}"
+      by (rule cSup_subset_mono) (auto simp: \<open>x\<in>s\<close>)
+    show "((\<lambda>k. Sup {f j x |j. k \<le> j}) \<longlongrightarrow> g x) sequentially"
+    proof (rule LIMSEQ_I, goal_cases)
+      case r: (1 r)
+      then have "0<r/2"
+        by auto
+      from assms(4)[THEN bspec, THEN LIMSEQ_D, OF x this] guess N .. note N=this
+      show ?case
+        apply (rule_tac x=N in exI,safe)
+        unfolding real_norm_def
+        apply (rule le_less_trans[of _ "r/2"])
+        apply (rule cSup_asclose)
+        apply safe
+        defer
+        apply (rule less_imp_le)
+        using N r
+        apply auto
+        done
+    qed
+  qed
+  note dec2 = conjunctD2[OF this]
+
+  show "g integrable_on s" by fact
+  show "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+  proof (rule LIMSEQ_I, goal_cases)
+    case r: (1 r)
+    from LIMSEQ_D [OF inc2(2) r] guess N1 .. note N1=this[unfolded real_norm_def]
+    from LIMSEQ_D [OF dec2(2) r] guess N2 .. note N2=this[unfolded real_norm_def]
+    show ?case
+    proof (rule_tac x="N1+N2" in exI, safe)
+      fix n
+      assume n: "n \<ge> N1 + N2"
+      have *: "\<And>i0 i i1 g. \<bar>i0 - g\<bar> < r \<longrightarrow> \<bar>i1 - g\<bar> < r \<longrightarrow> i0 \<le> i \<longrightarrow> i \<le> i1 \<longrightarrow> \<bar>i - g\<bar> < r"
+        by arith
+      show "norm (integral s (f n) - integral s g) < r"
+        unfolding real_norm_def
+      proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
+        show "integral s (\<lambda>x. Inf {f j x |j. n \<le> j}) \<le> integral s (f n)"
+          by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower)
+        show "integral s (f n) \<le> integral s (\<lambda>x. Sup {f j x |j. n \<le> j})"
+          by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper)
+      qed (insert n, auto)
+    qed
+  qed
+qed
+
+lemma dominated_convergence:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes f: "\<And>k. (f k) integrable_on s" and h: "h integrable_on s"
+    and le: "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
+    and conv: "\<forall>x \<in> s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
+  shows "g integrable_on s"
+    and "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+proof -
+  {
+    fix b :: 'm assume b: "b \<in> Basis"
+    have A: "(\<lambda>x. g x \<bullet> b) integrable_on s \<and>
+               (\<lambda>k. integral s (\<lambda>x. f k x \<bullet> b)) \<longlonglongrightarrow> integral s (\<lambda>x. g x \<bullet> b)"
+    proof (rule dominated_convergence_real)
+      fix k :: nat
+      from f show "(\<lambda>x. f k x \<bullet> b) integrable_on s" by (rule integrable_component)
+    next
+      from h show "h integrable_on s" .
+    next
+      fix k :: nat
+      show "\<forall>x\<in>s. norm (f k x \<bullet> b) \<le> h x"
+      proof
+        fix x assume x: "x \<in> s"
+        have "norm (f k x \<bullet> b) \<le> norm (f k x)" by (simp add: Basis_le_norm b)
+        also have "\<dots> \<le> h x" using le[of k] x by simp
+        finally show "norm (f k x \<bullet> b) \<le> h x" .
+      qed
+    next
+      from conv show "\<forall>x\<in>s. (\<lambda>k. f k x \<bullet> b) \<longlonglongrightarrow> g x \<bullet> b"
+        by (auto intro!: tendsto_inner)
+    qed
+    have B: "integral s ((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f k x \<bullet> b)) = integral s (\<lambda>x. f k x \<bullet> b) *\<^sub>R b"
+      for k by (rule integral_linear)
+               (simp_all add: f integrable_component bounded_linear_scaleR_left)
+    have C: "integral s ((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. g x \<bullet> b)) = integral s (\<lambda>x. g x \<bullet> b) *\<^sub>R b"
+      using A by (intro integral_linear)
+                 (simp_all add: integrable_component bounded_linear_scaleR_left)
+    note A B C
+  } note A = this
+
+  show "g integrable_on s" by (rule integrable_componentwise) (insert A, blast)
+  with A f show "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
+    by (subst (1 2) integral_componentwise)
+       (auto intro!: tendsto_setsum tendsto_scaleR simp: o_def)
+qed
+
+lemma has_integral_dominated_convergence:
+  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
+  assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s"
+    "\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
+    and x: "y \<longlonglongrightarrow> x"
+  shows "(g has_integral x) s"
+proof -
+  have int_f: "\<And>k. (f k) integrable_on s"
+    using assms by (auto simp: integrable_on_def)
+  have "(g has_integral (integral s g)) s"
+    by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+
+  moreover have "integral s g = x"
+  proof (rule LIMSEQ_unique)
+    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> x"
+      using integral_unique[OF assms(1)] x by simp
+    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> integral s g"
+      by (intro dominated_convergence[OF int_f assms(2)]) fact+
+  qed
+  ultimately show ?thesis
+    by simp
+qed
+
+end
+
+
+subsection \<open>Integration by parts\<close>
+
+lemma integration_by_parts_interior_strong:
+  assumes bilinear: "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)"
+  assumes s: "finite s" and le: "a \<le> b"
+  assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+                 "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes int: "((\<lambda>x. prod (f x) (g' x)) has_integral
+                  (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
+  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
+proof -
+  interpret bounded_bilinear prod by fact
+  have "((\<lambda>x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral
+          (prod (f b) (g b) - prod (f a) (g a))) {a..b}"
+    using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le])
+                   (auto intro!: continuous_intros continuous_on has_vector_derivative)
+  from has_integral_sub[OF this int] show ?thesis by (simp add: algebra_simps)
+qed
+
+lemma integration_by_parts_interior:
+  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
+          "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+          "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
+  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
+  by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+
+lemma integration_by_parts:
+  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
+          "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+          "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
+  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
+  by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all)
+
+lemma integrable_by_parts_interior_strong:
+  assumes bilinear: "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)"
+  assumes s: "finite s" and le: "a \<le> b"
+  assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+                 "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes int: "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
+  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
+proof -
+  from int obtain I where "((\<lambda>x. prod (f x) (g' x)) has_integral I) {a..b}"
+    unfolding integrable_on_def by blast
+  hence "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) -
+           (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp
+  from integration_by_parts_interior_strong[OF assms(1-7) this]
+    show ?thesis unfolding integrable_on_def by blast
+qed
+
+lemma integrable_by_parts_interior:
+  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
+          "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+          "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
+  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
+  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+
+lemma integrable_by_parts:
+  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
+          "continuous_on {a..b} f" "continuous_on {a..b} g"
+  assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
+          "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+  assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
+  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
+  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
+
+
+subsection \<open>Integration by substitution\<close>
+
+
+lemma has_integral_substitution_strong:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
+  assumes s: "finite s" and le: "a \<le> b" "c \<le> d" "g a \<le> g b"
+      and subset: "g ` {a..b} \<subseteq> {c..d}"
+      and f [continuous_intros]: "continuous_on {c..d} f"
+      and g [continuous_intros]: "continuous_on {a..b} g"
+      and deriv [derivative_intros]:
+              "\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
+    shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
+proof -
+  let ?F = "\<lambda>x. integral {c..g x} f"
+  have cont_int: "continuous_on {a..b} ?F"
+    by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous
+          f integrable_continuous_real)+
+  have deriv: "(((\<lambda>x. integral {c..x} f) \<circ> g) has_vector_derivative g' x *\<^sub>R f (g x))
+                 (at x within {a..b})" if "x \<in> {a..b} - s" for x
+    apply (rule has_vector_derivative_eq_rhs)
+    apply (rule vector_diff_chain_within)
+    apply (subst has_field_derivative_iff_has_vector_derivative [symmetric])
+    apply (rule deriv that)+
+    apply (rule has_vector_derivative_within_subset)
+    apply (rule integral_has_vector_derivative f)+
+    using that le subset
+    apply blast+
+    done
+  have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x))
+                  (at x)" if "x \<in> {a..b} - (s \<union> {a,b})" for x
+    using deriv[of x] that by (simp add: at_within_closed_interval o_def)
+
+
+  have "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}"
+    using le cont_int s deriv cont_int
+    by (intro fundamental_theorem_of_calculus_interior_strong[of "s \<union> {a,b}"]) simp_all
+  also from subset have "g x \<in> {c..d}" if "x \<in> {a..b}" for x using that by blast
+  from this[of a] this[of b] le have "c \<le> g a" "g b \<le> d" by auto
+  hence "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f"
+    by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all
+  hence "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f"
+    by (simp add: algebra_simps)
+  finally show ?thesis .
+qed
+
+lemma has_integral_substitution:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
+  assumes "a \<le> b" "c \<le> d" "g a \<le> g b" "g ` {a..b} \<subseteq> {c..d}"
+      and "continuous_on {c..d} f"
+      and "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
+    shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
+  by (intro has_integral_substitution_strong[of "{}" a b c d] assms)
+     (auto intro: DERIV_continuous_on assms)
+
+
+subsection \<open>Compute a double integral using iterated integrals and switching the order of integration\<close>
+
+lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
+  by auto
+
+lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
+  by auto
+
+lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
+  by blast
+
+lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
+  by blast
+
+lemma continuous_on_imp_integrable_on_Pair1:
+  fixes f :: "_ \<Rightarrow> 'b::banach"
+  assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \<in> cbox a b"
+  shows "(\<lambda>y. f (x, y)) integrable_on (cbox c d)"
+proof -
+  have "f \<circ> (\<lambda>y. (x, y)) integrable_on (cbox c d)"
+    apply (rule integrable_continuous)
+    apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]])
+    using x
+    apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con)
+    done
+  then show ?thesis
+    by (simp add: o_def)
+qed
+
+lemma integral_integrable_2dim:
+  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
+  assumes "continuous_on (cbox (a,c) (b,d)) f"
+    shows "(\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y))) integrable_on cbox a b"
+proof (cases "content(cbox c d) = 0")
+case True
+  then show ?thesis
+    by (simp add: True integrable_const)
+next
+  case False
+  have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f"
+    by (simp add: assms compact_cbox compact_uniformly_continuous)
+  { fix x::'a and e::real
+    assume x: "x \<in> cbox a b" and e: "0 < e"
+    then have e2_gt: "0 < e / 2 / content (cbox c d)" and e2_less: "e / 2 / content (cbox c d) * content (cbox c d) < e"
+      by (auto simp: False content_lt_nz e)
+    then obtain dd
+    where dd: "\<And>x x'. \<lbrakk>x\<in>cbox (a, c) (b, d); x'\<in>cbox (a, c) (b, d); norm (x' - x) < dd\<rbrakk>
+                       \<Longrightarrow> norm (f x' - f x) \<le> e / (2 * content (cbox c d))"  "dd>0"
+      using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e / (2 * content (cbox c d))"]
+      by (auto simp: dist_norm intro: less_imp_le)
+    have "\<exists>delta>0. \<forall>x'\<in>cbox a b. norm (x' - x) < delta \<longrightarrow> norm (integral (cbox c d) (\<lambda>u. f (x', u) - f (x, u))) < e"
+      apply (rule_tac x=dd in exI)
+      using dd e2_gt assms x
+      apply clarify
+      apply (rule le_less_trans [OF _ e2_less])
+      apply (rule integrable_bound)
+      apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1)
+      done
+  } note * = this
+  show ?thesis
+    apply (rule integrable_continuous)
+    apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms])
+    done
+qed
+
+lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk>
+            \<Longrightarrow> norm(y - x) \<le> e"
+using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"]
+  by (simp add: add_diff_add)
+
+lemma integral_split:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
+  assumes f: "f integrable_on (cbox a b)"
+      and k: "k \<in> Basis"
+  shows "integral (cbox a b) f =
+           integral (cbox a b \<inter> {x. x\<bullet>k \<le> c}) f +
+           integral (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) f"
+  apply (rule integral_unique [OF has_integral_split [where c=c]])
+  using k f
+  apply (auto simp: has_integral_integral [symmetric])
+  done
+
+lemma integral_swap_operative:
+  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
+  assumes f: "continuous_on s f" and e: "0 < e"
+    shows "comm_monoid.operative (op \<and>) True
+           (\<lambda>k. \<forall>a b c d.
+                cbox (a,c) (b,d) \<subseteq> k \<and> cbox (a,c) (b,d) \<subseteq> s
+                \<longrightarrow> norm(integral (cbox (a,c) (b,d)) f -
+                         integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f((x,y)))))
+                    \<le> e * content (cbox (a,c) (b,d)))"
+proof (auto simp: comm_monoid.operative_def[OF comm_monoid_and])
+  fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b
+  assume c0: "content (cbox (a, c) (b, d)) = 0"
+     and cb1: "cbox (u, w) (v, z) \<subseteq> cbox (a, c) (b, d)"
+     and cb2: "cbox (u, w) (v, z) \<subseteq> s"
+  have c0': "content (cbox (u, w) (v, z)) = 0"
+    by (fact content_0_subset [OF c0 cb1])
+  show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
+          \<le> e * content (cbox (u,w) (v,z))"
+    using content_cbox_pair_eq0_D [OF c0']
+    by (force simp add: c0')
+next
+  fix a::'a and c::'b and b::'a and d::'b
+  and M::real and i::'a and j::'b
+  and u::'a and v::'a and w::'b and z::'b
+  assume ij: "(i,j) \<in> Basis"
+     and n1: "\<forall>a' b' c' d'.
+                cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
+                cbox (a',c') (b',d') \<subseteq> {x. x \<bullet> (i,j) \<le> M} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
+                norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
+                \<le> e * content (cbox (a',c') (b',d'))"
+     and n2: "\<forall>a' b' c' d'.
+                cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
+                cbox (a',c') (b',d') \<subseteq> {x. M \<le> x \<bullet> (i,j)} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
+                norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
+                \<le> e * content (cbox (a',c') (b',d'))"
+     and subs: "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)"  "cbox (u,w) (v,z) \<subseteq> s"
+  have fcont: "continuous_on (cbox (u, w) (v, z)) f"
+    using assms(1) continuous_on_subset  subs(2) by blast
+  then have fint: "f integrable_on cbox (u, w) (v, z)"
+    by (metis integrable_continuous)
+  consider "i \<in> Basis" "j=0" | "j \<in> Basis" "i=0"  using ij
+    by (auto simp: Euclidean_Space.Basis_prod_def)
+  then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
+             \<le> e * content (cbox (u,w) (v,z))" (is ?normle)
+  proof cases
+    case 1
+    then have e: "e * content (cbox (u, w) (v, z)) =
+                  e * (content (cbox u v \<inter> {x. x \<bullet> i \<le> M}) * content (cbox w z)) +
+                  e * (content (cbox u v \<inter> {x. M \<le> x \<bullet> i}) * content (cbox w z))"
+      by (simp add: content_split [where c=M] content_Pair algebra_simps)
+    have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
+                integral (cbox u v \<inter> {x. x \<bullet> i \<le> M}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) +
+                integral (cbox u v \<inter> {x. M \<le> x \<bullet> i}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))"
+      using 1 f subs integral_integrable_2dim continuous_on_subset
+      by (blast intro: integral_split)
+    show ?normle
+      apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
+      using 1 subs
+      apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\<lambda>u. M\<le>u"] setcomp_dot1 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp1)
+      apply (simp_all add: interval_split ij)
+      apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
+      apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
+      apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
+      done
+  next
+    case 2
+    then have e: "e * content (cbox (u, w) (v, z)) =
+                  e * (content (cbox u v) * content (cbox w z \<inter> {x. x \<bullet> j \<le> M})) +
+                  e * (content (cbox u v) * content (cbox w z \<inter> {x. M \<le> x \<bullet> j}))"
+      by (simp add: content_split [where c=M] content_Pair algebra_simps)
+    have "(\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
+                "(\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
+      using 2 subs
+      apply (simp_all add: interval_split)
+      apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]])
+      apply (auto simp: cbox_Pair_eq interval_split [symmetric])
+      done
+    with 2 have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
+                   integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) +
+                   integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y)))"
+      by (simp add: integral_add [symmetric] integral_split [symmetric]
+                    continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong)
+    show ?normle
+      apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
+      using 2 subs
+      apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\<lambda>u. M\<le>u"] setcomp_dot2 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp2)
+      apply (simp_all add: interval_split ij)
+      apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
+      apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
+      apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
+      done
+  qed
+qed
+
+lemma integral_Pair_const:
+    "integral (cbox (a,c) (b,d)) (\<lambda>x. k) =
+     integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. k))"
+  by (simp add: content_Pair)
+
+lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)"
+  by (simp add: norm_minus_eqI)
+
+lemma integral_prod_continuous:
+  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
+  assumes "continuous_on (cbox (a,c) (b,d)) f" (is "continuous_on ?CBOX f")
+    shows "integral (cbox (a,c) (b,d)) f = integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f(x,y)))"
+proof (cases "content ?CBOX = 0")
+  case True
+  then show ?thesis
+    by (auto simp: content_Pair)
+next
+  case False
+  then have cbp: "content ?CBOX > 0"
+    using content_lt_nz by blast
+  have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) = 0"
+  proof (rule dense_eq0_I, simp)
+    fix e::real  assume "0 < e"
+    with cbp have e': "0 < e / content ?CBOX"
+      by simp
+    have f_int_acbd: "f integrable_on cbox (a,c) (b,d)"
+      by (rule integrable_continuous [OF assms])
+    { fix p
+      assume p: "p division_of cbox (a,c) (b,d)"
+      note opd1 = comm_monoid_set.operative_division [OF comm_monoid_set_and integral_swap_operative [OF assms e'], THEN iffD1,
+                       THEN spec, THEN spec, THEN spec, THEN spec, of p "(a,c)" "(b,d)" a c b d]
+      have "(\<And>t u v w z.
+              \<lbrakk>t \<in> p; cbox (u,w) (v,z) \<subseteq> t;  cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)\<rbrakk> \<Longrightarrow>
+              norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
+              \<le> e * content (cbox (u,w) (v,z)) / content?CBOX)
+            \<Longrightarrow>
+            norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
+        using opd1 [OF p] False  by (simp add: comm_monoid_set_F_and)
+    } note op_acbd = this
+    { fix k::real and p and u::'a and v w and z::'b and t1 t2 l
+      assume k: "0 < k"
+         and nf: "\<And>x y u v.
+                  \<lbrakk>x \<in> cbox a b; y \<in> cbox c d; u \<in> cbox a b; v\<in>cbox c d; norm (u-x, v-y) < k\<rbrakk>
+                  \<Longrightarrow> norm (f(u,v) - f(x,y)) < e / (2 * (content ?CBOX))"
+         and p_acbd: "p tagged_division_of cbox (a,c) (b,d)"
+         and fine: "(\<lambda>x. ball x k) fine p"  "((t1,t2), l) \<in> p"
+         and uwvz_sub: "cbox (u,w) (v,z) \<subseteq> l" "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)"
+      have t: "t1 \<in> cbox a b" "t2 \<in> cbox c d"
+        by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+
+      have ls: "l \<subseteq> ball (t1,t2) k"
+        using fine by (simp add: fine_def Ball_def)
+      { fix x1 x2
+        assume xuvwz: "x1 \<in> cbox u v" "x2 \<in> cbox w z"
+        then have x: "x1 \<in> cbox a b" "x2 \<in> cbox c d"
+          using uwvz_sub by auto
+        have "norm (x1 - t1, x2 - t2) < k"
+          using xuvwz ls uwvz_sub unfolding ball_def
+          by (force simp add: cbox_Pair_eq dist_norm norm_minus2)
+        then have "norm (f (x1,x2) - f (t1,t2)) \<le> e / content ?CBOX / 2"
+          using nf [OF t x]  by simp
+      } note nf' = this
+      have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)"
+        using f_int_acbd uwvz_sub integrable_on_subcbox by blast
+      have f_int_uv: "\<And>x. x \<in> cbox u v \<Longrightarrow> (\<lambda>y. f (x,y)) integrable_on cbox w z"
+        using assms continuous_on_subset uwvz_sub
+        by (blast intro: continuous_on_imp_integrable_on_Pair1)
+      have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\<lambda>x. f (t1,t2)))
+         \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
+        apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const)
+        apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]])
+        using cbp e' nf'
+        apply (auto simp: integrable_diff f_int_uwvz integrable_const)
+        done
+      have int_integrable: "(\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) integrable_on cbox u v"
+        using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast
+      have normint_wz:
+         "\<And>x. x \<in> cbox u v \<Longrightarrow>
+               norm (integral (cbox w z) (\<lambda>y. f (x, y)) - integral (cbox w z) (\<lambda>y. f (t1, t2)))
+               \<le> e * content (cbox w z) / content (cbox (a, c) (b, d)) / 2"
+        apply (simp only: integral_diff [symmetric] f_int_uv integrable_const)
+        apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]])
+        using cbp e' nf'
+        apply (auto simp: integrable_diff f_int_uv integrable_const)
+        done
+      have "norm (integral (cbox u v)
+               (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)) - integral (cbox w z) (\<lambda>y. f (t1,t2))))
+            \<le> e * content (cbox w z) / content ?CBOX / 2 * content (cbox u v)"
+        apply (rule integrable_bound [OF _ _ normint_wz])
+        using cbp e'
+        apply (auto simp: divide_simps content_pos_le integrable_diff int_integrable integrable_const)
+        done
+      also have "... \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
+        by (simp add: content_Pair divide_simps)
+      finally have 2: "norm (integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))) -
+                      integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (t1,t2))))
+                \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
+        by (simp only: integral_diff [symmetric] int_integrable integrable_const)
+      have norm_xx: "\<lbrakk>x' = y'; norm(x - x') \<le> e/2; norm(y - y') \<le> e/2\<rbrakk> \<Longrightarrow> norm(x - y) \<le> e" for x::'c and y x' y' e
+        using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] real_sum_of_halves
+        by (simp add: norm_minus_commute)
+      have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
+            \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX"
+        by (rule norm_xx [OF integral_Pair_const 1 2])
+    } note * = this
+    show "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
+      using compact_uniformly_continuous [OF assms compact_cbox]
+      apply (simp add: uniformly_continuous_on_def dist_norm)
+      apply (drule_tac x="e / 2 / content?CBOX" in spec)
+      using cbp \<open>0 < e\<close>
+      apply (auto simp: zero_less_mult_iff)
+      apply (rename_tac k)
+      apply (rule_tac e1=k in fine_division_exists [OF gauge_ball, where a = "(a,c)" and b = "(b,d)"])
+      apply assumption
+      apply (rule op_acbd)
+      apply (erule division_of_tagged_division)
+      using *
+      apply auto
+      done
+  qed
+  then show ?thesis
+    by simp
+qed
+
+lemma swap_continuous:
+  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
+    shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)"
+proof -
+  have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap"
+    by auto
+  then show ?thesis
+    apply (rule ssubst)
+    apply (rule continuous_on_compose)
+    apply (simp add: split_def)
+    apply (rule continuous_intros | simp add: assms)+
+    done
+qed
+
+lemma integral_swap_2dim:
+  fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
+  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
+    shows "integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y) = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
+proof -
+  have "((\<lambda>(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))"
+    apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\<lambda>(x,y). f y x" "integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)", simplified])
+    apply (auto simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)
+    done
+ then show ?thesis
+   by force
+qed
+
+theorem integral_swap_continuous:
+  fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
+  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
+    shows "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) =
+           integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
+proof -
+  have "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y)"
+    using integral_prod_continuous [OF assms] by auto
+  also have "... = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
+    by (rule integral_swap_2dim [OF assms])
+  also have "... = integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
+    using integral_prod_continuous [OF swap_continuous] assms
+    by auto
+  finally show ?thesis .
+qed
+
+
+subsection \<open>Definite integrals for exponential and power function\<close>
+
+lemma has_integral_exp_minus_to_infinity:
+  assumes a: "a > 0"
+  shows   "((\<lambda>x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}"
+proof -
+  define f where "f = (\<lambda>k x. if x \<in> {c..real k} then exp (-a*x) else 0)"
+
+  {
+    fix k :: nat assume k: "of_nat k \<ge> c"
+    from k a
+      have "((\<lambda>x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}"
+      by (intro fundamental_theorem_of_calculus)
+         (auto intro!: derivative_eq_intros
+               simp: has_field_derivative_iff_has_vector_derivative [symmetric])
+    hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def
+      by (subst has_integral_restrict) simp_all
+  } note has_integral_f = this
+
+  have [simp]: "f k = (\<lambda>_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def)
+  have integral_f: "integral {c..} (f k) =
+                      (if real k \<ge> c then exp (-a*c)/a - exp (-a*real k)/a else 0)"
+    for k using integral_unique[OF has_integral_f[of k]] by simp
+
+  have A: "(\<lambda>x. exp (-a*x)) integrable_on {c..} \<and>
+             (\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> integral {c..} (\<lambda>x. exp (-a*x))"
+  proof (intro monotone_convergence_increasing allI ballI)
+    fix k ::nat
+    have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P)
+      unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real)
+    hence int: "(f k) integrable_on {c..of_real k}"
+      by (rule integrable_eq[rotated]) (simp add: f_def)
+    show "f k integrable_on {c..}"
+      by (rule integrable_on_superset[OF _ _ int]) (auto simp: f_def)
+  next
+    fix x assume x: "x \<in> {c..}"
+    have "sequentially \<le> principal {nat \<lceil>x\<rceil>..}" unfolding at_top_def by (simp add: Inf_lower)
+    also have "{nat \<lceil>x\<rceil>..} \<subseteq> {k. x \<le> real k}" by auto
+    also have "inf (principal \<dots>) (principal {k. \<not>x \<le> real k}) =
+                 principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" by simp
+    also have "{k. x \<le> real k} \<inter> {k. \<not>x \<le> real k} = {}" by blast
+    finally have "inf sequentially (principal {k. \<not>x \<le> real k}) = bot"
+      by (simp add: inf.coboundedI1 bot_unique)
+    with x show "(\<lambda>k. f k x) \<longlonglongrightarrow> exp (-a*x)" unfolding f_def
+      by (intro filterlim_If) auto
+  next
+    have "\<bar>integral {c..} (f k)\<bar> \<le> exp (-a*c)/a" for k
+    proof (cases "c > of_nat k")
+      case False
+      hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)"
+        by (simp add: integral_f)
+      also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) =
+                   exp (- (a * c)) / a - exp (- (a * real k)) / a"
+        using False a by (intro abs_of_nonneg) (simp_all add: field_simps)
+      also have "\<dots> \<le> exp (- a * c) / a" using a by simp
+      finally show ?thesis .
+    qed (insert a, simp_all add: integral_f)
+    thus "bounded {integral {c..} (f k) |k. True}"
+      by (intro bounded_realI[of _ "exp (-a*c)/a"]) auto
+  qed (auto simp: f_def)
+
+  from eventually_gt_at_top[of "nat \<lceil>c\<rceil>"] have "eventually (\<lambda>k. of_nat k > c) sequentially"
+    by eventually_elim linarith
+  hence "eventually (\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially"
+    by eventually_elim (simp add: integral_f)
+  moreover have "(\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
+    by (intro tendsto_intros filterlim_compose[OF exp_at_bot]
+          filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+
+       (insert a, simp_all)
+  ultimately have "(\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
+    by (rule Lim_transform_eventually)
+  from LIMSEQ_unique[OF conjunct2[OF A] this]
+    have "integral {c..} (\<lambda>x. exp (-a*x)) = exp (-a*c)/a" by simp
+  with conjunct1[OF A] show ?thesis
+    by (simp add: has_integral_integral)
+qed
+
+lemma integrable_on_exp_minus_to_infinity: "a > 0 \<Longrightarrow> (\<lambda>x. exp (-a*x) :: real) integrable_on {c..}"
+  using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast
+
+lemma has_integral_powr_from_0:
+  assumes a: "a > (-1::real)" and c: "c \<ge> 0"
+  shows   "((\<lambda>x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}"
+proof (cases "c = 0")
+  case False
+  define f where "f = (\<lambda>k x. if x \<in> {inverse (of_nat (Suc k))..c} then x powr a else 0)"
+  define F where "F = (\<lambda>k. if inverse (of_nat (Suc k)) \<le> c then
+                             c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)"
+
+  {
+    fix k :: nat
+    have "(f k has_integral F k) {0..c}"
+    proof (cases "inverse (of_nat (Suc k)) \<le> c")
+      case True
+      {
+        fix x assume x: "x \<ge> inverse (1 + real k)"
+        have "0 < inverse (1 + real k)" by simp
+        also note x
+        finally have "x > 0" .
+      } note x = this
+      hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) -
+               inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}"
+        using True a by (intro fundamental_theorem_of_calculus)
+           (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const
+             continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric])
+      with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all
+    next
+      case False
+      thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto
+    qed
+  } note has_integral_f = this
+  have integral_f: "integral {0..c} (f k) = F k" for k
+    using has_integral_f[of k] by (rule integral_unique)
+
+  have A: "(\<lambda>x. x powr a) integrable_on {0..c} \<and>
+           (\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> integral {0..c} (\<lambda>x. x powr a)"
+  proof (intro monotone_convergence_increasing ballI allI)
+    fix k from has_integral_f[of k] show "f k integrable_on {0..c}"
+      by (auto simp: integrable_on_def)
+  next
+    fix k :: nat and x :: real
+    {
+      assume x: "inverse (real (Suc k)) \<le> x"
+      have "inverse (real (Suc (Suc k))) \<le> inverse (real (Suc k))" by (simp add: field_simps)
+      also note x
+      finally have "inverse (real (Suc (Suc k))) \<le> x" .
+    }
+    thus "f k x \<le> f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc)
+  next
+    fix x assume x: "x \<in> {0..c}"
+    show "(\<lambda>k. f k x) \<longlonglongrightarrow> x powr a"
+    proof (cases "x = 0")
+      case False
+      with x have "x > 0" by simp
+      from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
+        have "eventually (\<lambda>k. x powr a = f k x) sequentially"
+        by eventually_elim (insert x, simp add: f_def)
+      moreover have "(\<lambda>_. x powr a) \<longlonglongrightarrow> x powr a" by simp
+      ultimately show ?thesis by (rule Lim_transform_eventually)
+    qed (simp_all add: f_def)
+  next
+    {
+      fix k
+      from a have "F k \<le> c powr (a + 1) / (a + 1)"
+        by (auto simp add: F_def divide_simps)
+      also from a have "F k \<ge> 0"
+        by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2)
+      hence "F k = abs (F k)" by simp
+      finally have "abs (F k) \<le>  c powr (a + 1) / (a + 1)" .
+    }
+    thus "bounded {integral {0..c} (f k) |k. True}"
+      by (intro bounded_realI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f)
+  qed
+
+  from False c have "c > 0" by simp
+  from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
+    have "eventually (\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) =
+            integral {0..c} (f k)) sequentially"
+    by eventually_elim (simp add: integral_f F_def)
+  moreover have "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
+                   \<longlonglongrightarrow> c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)"
+    using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto
+  hence "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
+          \<longlonglongrightarrow> c powr (a + 1) / (a + 1)" by simp
+  ultimately have "(\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> c powr (a+1) / (a+1)"
+    by (rule Lim_transform_eventually)
+  with A have "integral {0..c} (\<lambda>x. x powr a) = c powr (a+1) / (a+1)"
+    by (blast intro: LIMSEQ_unique)
+  with A show ?thesis by (simp add: has_integral_integral)
+qed (simp_all add: has_integral_refl)
+
+lemma integrable_on_powr_from_0:
+  assumes a: "a > (-1::real)" and c: "c \<ge> 0"
+  shows   "(\<lambda>x. x powr a) integrable_on {0..c}"
+  using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast
+
+end
--- a/src/HOL/Multivariate_Analysis/Integral_Test.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Integral_Test.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -1,18 +1,18 @@
 (*  Title:    HOL/Multivariate_Analysis/Integral_Test.thy
     Author:   Manuel Eberl, TU München
 *)
-  
+
 section \<open>Integral Test for Summability\<close>
 
 theory Integral_Test
-imports Integration
+imports Henstock_Kurzweil_Integration
 begin
 
 text \<open>
-  The integral test for summability. We show here that for a decreasing non-negative 
-  function, the infinite sum over that function evaluated at the natural numbers 
+  The integral test for summability. We show here that for a decreasing non-negative
+  function, the infinite sum over that function evaluated at the natural numbers
   converges iff the corresponding integral converges.
-  
+
   As a useful side result, we also provide some results on the difference between
   the integral and the partial sum. (This is useful e.g. for the definition of the
   Euler-Mascheroni constant)
@@ -33,7 +33,7 @@
 proof -
   note int = integrable_continuous_real[OF continuous_on_subset[OF cont]]
   let ?int = "\<lambda>a b. integral {of_nat a..of_nat b} f"
-  have "-sum_integral_diff_series n = ?int 0 n - (\<Sum>k\<le>n. f (of_nat k))" 
+  have "-sum_integral_diff_series n = ?int 0 n - (\<Sum>k\<le>n. f (of_nat k))"
     by (simp add: sum_integral_diff_series_def)
   also have "?int 0 n = (\<Sum>k<n. ?int k (Suc k))"
   proof (induction n)
@@ -60,7 +60,7 @@
   have d_mono: "sum_integral_diff_series (Suc n) \<le> sum_integral_diff_series n" for n
   proof -
     fix n :: nat
-    have "sum_integral_diff_series (Suc n) - sum_integral_diff_series n = 
+    have "sum_integral_diff_series (Suc n) - sum_integral_diff_series n =
             f (of_nat (Suc n)) + (?int 0 n - ?int 0 (Suc n))"
       unfolding sum_integral_diff_series_def by (simp add: algebra_simps)
     also have "?int 0 n - ?int 0 (Suc n) = -?int n (Suc n)"
@@ -77,7 +77,7 @@
 
 lemma sum_integral_diff_series_Bseq: "Bseq sum_integral_diff_series"
 proof -
-  from sum_integral_diff_series_nonneg and sum_integral_diff_series_antimono 
+  from sum_integral_diff_series_nonneg and sum_integral_diff_series_antimono
     have "norm (sum_integral_diff_series n) \<le> sum_integral_diff_series 0" for n by simp
   thus "Bseq sum_integral_diff_series" by (rule BseqI')
 qed
@@ -97,12 +97,12 @@
   also have "... \<longleftrightarrow> convergent (\<lambda>n. integral {0..of_nat n} f)"
   proof
     assume "convergent (\<lambda>n. \<Sum>k\<le>n. f (of_nat k))"
-    from convergent_diff[OF this sum_integral_diff_series_convergent] 
-      show "convergent (\<lambda>n. integral {0..of_nat n} f)" 
+    from convergent_diff[OF this sum_integral_diff_series_convergent]
+      show "convergent (\<lambda>n. integral {0..of_nat n} f)"
         unfolding sum_integral_diff_series_def by simp
   next
     assume "convergent (\<lambda>n. integral {0..of_nat n} f)"
-    from convergent_add[OF this sum_integral_diff_series_convergent] 
+    from convergent_add[OF this sum_integral_diff_series_convergent]
       show "convergent (\<lambda>n. \<Sum>k\<le>n. f (of_nat k))" unfolding sum_integral_diff_series_def by simp
   qed
   finally show ?thesis by simp
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,12645 +0,0 @@
-(*  Author:     John Harrison
-    Author:     Robert Himmelmann, TU Muenchen (Translation from HOL light); proofs reworked by LCP
-*)
-
-section \<open>Kurzweil-Henstock Gauge Integration in many dimensions.\<close>
-
-theory Integration
-imports
-  Derivative
-  Uniform_Limit
-  "~~/src/HOL/Library/Indicator_Function"
-begin
-
-lemmas scaleR_simps = scaleR_zero_left scaleR_minus_left scaleR_left_diff_distrib
-  scaleR_zero_right scaleR_minus_right scaleR_right_diff_distrib scaleR_eq_0_iff
-  scaleR_cancel_left scaleR_cancel_right scaleR_add_right scaleR_add_left real_vector_class.scaleR_one
-
-
-subsection \<open>Sundries\<close>
-
-lemma conjunctD2: assumes "a \<and> b" shows a b using assms by auto
-lemma conjunctD3: assumes "a \<and> b \<and> c" shows a b c using assms by auto
-lemma conjunctD4: assumes "a \<and> b \<and> c \<and> d" shows a b c d using assms by auto
-
-declare norm_triangle_ineq4[intro]
-
-lemma simple_image: "{f x |x . x \<in> s} = f ` s"
-  by blast
-
-lemma linear_simps:
-  assumes "bounded_linear f"
-  shows
-    "f (a + b) = f a + f b"
-    "f (a - b) = f a - f b"
-    "f 0 = 0"
-    "f (- a) = - f a"
-    "f (s *\<^sub>R v) = s *\<^sub>R (f v)"
-proof -
-  interpret f: bounded_linear f by fact
-  show "f (a + b) = f a + f b" by (rule f.add)
-  show "f (a - b) = f a - f b" by (rule f.diff)
-  show "f 0 = 0" by (rule f.zero)
-  show "f (- a) = - f a" by (rule f.minus)
-  show "f (s *\<^sub>R v) = s *\<^sub>R (f v)" by (rule f.scaleR)
-qed
-
-lemma bounded_linearI:
-  assumes "\<And>x y. f (x + y) = f x + f y"
-    and "\<And>r x. f (r *\<^sub>R x) = r *\<^sub>R f x"
-    and "\<And>x. norm (f x) \<le> norm x * K"
-  shows "bounded_linear f"
-  using assms by (rule bounded_linear_intro) (* FIXME: duplicate *)
-
-lemma bounded_linear_component [intro]: "bounded_linear (\<lambda>x::'a::euclidean_space. x \<bullet> k)"
-  by (rule bounded_linear_inner_left)
-
-lemma transitive_stepwise_lt_eq:
-  assumes "(\<And>x y z::nat. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z)"
-  shows "((\<forall>m. \<forall>n>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n)))"
-  (is "?l = ?r")
-proof safe
-  assume ?r
-  fix n m :: nat
-  assume "m < n"
-  then show "R m n"
-  proof (induct n arbitrary: m)
-    case 0
-    then show ?case by auto
-  next
-    case (Suc n)
-    show ?case
-    proof (cases "m < n")
-      case True
-      show ?thesis
-        apply (rule assms[OF Suc(1)[OF True]])
-        using \<open>?r\<close>
-        apply auto
-        done
-    next
-      case False
-      then have "m = n"
-        using Suc(2) by auto
-      then show ?thesis
-        using \<open>?r\<close> by auto
-    qed
-  qed
-qed auto
-
-lemma transitive_stepwise_gt:
-  assumes "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z" "\<And>n. R n (Suc n)"
-  shows "\<forall>n>m. R m n"
-proof -
-  have "\<forall>m. \<forall>n>m. R m n"
-    apply (subst transitive_stepwise_lt_eq)
-    apply (blast intro: assms)+
-    done
-  then show ?thesis by auto
-qed
-
-lemma transitive_stepwise_le_eq:
-  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
-  shows "(\<forall>m. \<forall>n\<ge>m. R m n) \<longleftrightarrow> (\<forall>n. R n (Suc n))"
-  (is "?l = ?r")
-proof safe
-  assume ?r
-  fix m n :: nat
-  assume "m \<le> n"
-  then show "R m n"
-  proof (induct n arbitrary: m)
-    case 0
-    with assms show ?case by auto
-  next
-    case (Suc n)
-    show ?case
-    proof (cases "m \<le> n")
-      case True
-      with Suc.hyps \<open>\<forall>n. R n (Suc n)\<close> assms show ?thesis
-        by blast
-    next
-      case False
-      then have "m = Suc n"
-        using Suc(2) by auto
-      then show ?thesis
-        using assms(1) by auto
-    qed
-  qed
-qed auto
-
-lemma transitive_stepwise_le:
-  assumes "\<And>x. R x x" "\<And>x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
-    and "\<And>n. R n (Suc n)"
-  shows "\<forall>n\<ge>m. R m n"
-proof -
-  have "\<forall>m. \<forall>n\<ge>m. R m n"
-    apply (subst transitive_stepwise_le_eq)
-    apply (blast intro: assms)+
-    done
-  then show ?thesis by auto
-qed
-
-
-subsection \<open>Some useful lemmas about intervals.\<close>
-
-lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)"
-  using nonempty_Basis
-  by (fastforce simp add: set_eq_iff mem_box)
-
-lemma interior_subset_union_intervals:
-  assumes "i = cbox a b"
-    and "j = cbox c d"
-    and "interior j \<noteq> {}"
-    and "i \<subseteq> j \<union> s"
-    and "interior i \<inter> interior j = {}"
-  shows "interior i \<subseteq> interior s"
-proof -
-  have "box a b \<inter> cbox c d = {}"
-     using inter_interval_mixed_eq_empty[of c d a b] and assms(3,5)
-     unfolding assms(1,2) interior_cbox by auto
-  moreover
-  have "box a b \<subseteq> cbox c d \<union> s"
-    apply (rule order_trans,rule box_subset_cbox)
-    using assms(4) unfolding assms(1,2)
-    apply auto
-    done
-  ultimately
-  show ?thesis
-    unfolding assms interior_cbox
-      by auto (metis IntI UnE empty_iff interior_maximal open_box subsetCE subsetI)
-qed
-
-lemma inter_interior_unions_intervals:
-  fixes f::"('a::euclidean_space) set set"
-  assumes "finite f"
-    and "open s"
-    and "\<forall>t\<in>f. \<exists>a b. t = cbox a b"
-    and "\<forall>t\<in>f. s \<inter> (interior t) = {}"
-  shows "s \<inter> interior (\<Union>f) = {}"
-proof (clarsimp simp only: all_not_in_conv [symmetric])
-  fix x
-  assume x: "x \<in> s" "x \<in> interior (\<Union>f)"
-  have lem1: "\<And>x e s U. ball x e \<subseteq> s \<inter> interior U \<longleftrightarrow> ball x e \<subseteq> s \<inter> U"
-    using interior_subset
-    by auto (meson Topology_Euclidean_Space.open_ball contra_subsetD interior_maximal mem_ball)
-  have "\<exists>t\<in>f. \<exists>x. \<exists>e>0. ball x e \<subseteq> s \<inter> t"
-    if "finite f" and "\<forall>t\<in>f. \<exists>a b. t = cbox a b" and "\<exists>x. x \<in> s \<inter> interior (\<Union>f)" for f
-    using that
-  proof (induct rule: finite_induct)
-    case empty
-    obtain x where "x \<in> s \<inter> interior (\<Union>{})"
-      using empty(2) ..
-    then have False
-      unfolding Union_empty interior_empty by auto
-    then show ?case by auto
-  next
-    case (insert i f)
-    obtain x where x: "x \<in> s \<inter> interior (\<Union>insert i f)"
-      using insert(5) ..
-    then obtain e where e: "0 < e \<and> ball x e \<subseteq> s \<inter> interior (\<Union>insert i f)"
-      unfolding open_contains_ball_eq[OF open_Int[OF assms(2) open_interior], rule_format] ..
-    obtain a where "\<exists>b. i = cbox a b"
-      using insert(4)[rule_format,OF insertI1] ..
-    then obtain b where ab: "i = cbox a b" ..
-    show ?case
-    proof (cases "x \<in> i")
-      case False
-      then have "x \<in> UNIV - cbox a b"
-        unfolding ab by auto
-      then obtain d where "0 < d \<and> ball x d \<subseteq> UNIV - cbox a b"
-        unfolding open_contains_ball_eq[OF open_Diff[OF open_UNIV closed_cbox],rule_format] ..
-      then have "0 < d" "ball x (min d e) \<subseteq> UNIV - i"
-        unfolding ab ball_min_Int by auto
-      then have "ball x (min d e) \<subseteq> s \<inter> interior (\<Union>f)"
-        using e unfolding lem1 unfolding  ball_min_Int by auto
-      then have "x \<in> s \<inter> interior (\<Union>f)" using \<open>d>0\<close> e by auto
-      then have "\<exists>t\<in>f. \<exists>x e. 0 < e \<and> ball x e \<subseteq> s \<inter> t"
-        using insert.hyps(3) insert.prems(1) by blast
-      then show ?thesis by auto
-    next
-      case True show ?thesis
-      proof (cases "x\<in>box a b")
-        case True
-        then obtain d where "0 < d \<and> ball x d \<subseteq> box a b"
-          unfolding open_contains_ball_eq[OF open_box,rule_format] ..
-        then show ?thesis
-          apply (rule_tac x=i in bexI, rule_tac x=x in exI, rule_tac x="min d e" in exI)
-          unfolding ab
-          using box_subset_cbox[of a b] and e
-          apply fastforce+
-          done
-      next
-        case False
-        then obtain k where "x\<bullet>k \<le> a\<bullet>k \<or> x\<bullet>k \<ge> b\<bullet>k" and k: "k \<in> Basis"
-          unfolding mem_box by (auto simp add: not_less)
-        then have "x\<bullet>k = a\<bullet>k \<or> x\<bullet>k = b\<bullet>k"
-          using True unfolding ab and mem_box
-            apply (erule_tac x = k in ballE)
-            apply auto
-            done
-        then have "\<exists>x. ball x (e/2) \<subseteq> s \<inter> (\<Union>f)"
-        proof (rule disjE)
-          let ?z = "x - (e/2) *\<^sub>R k"
-          assume as: "x\<bullet>k = a\<bullet>k"
-          have "ball ?z (e / 2) \<inter> i = {}"
-          proof (clarsimp simp only: all_not_in_conv [symmetric])
-            fix y
-            assume "y \<in> ball ?z (e / 2)" and yi: "y \<in> i"
-            then have "dist ?z y < e/2" by auto
-            then have "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
-              using Basis_le_norm[OF k, of "?z - y"] unfolding dist_norm by auto
-            then have "y\<bullet>k < a\<bullet>k"
-              using e k
-              by (auto simp add: field_simps abs_less_iff as inner_simps)
-            then have "y \<notin> i"
-              unfolding ab mem_box by (auto intro!: bexI[OF _ k])
-            then show False using yi by auto
-          qed
-          moreover
-          have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)"
-            apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]])
-          proof
-            fix y
-            assume as: "y \<in> ball ?z (e/2)"
-            have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y - (e / 2) *\<^sub>R k)"
-              apply (rule order_trans,rule norm_triangle_sub[of "x - y" "(e/2) *\<^sub>R k"])
-              unfolding norm_scaleR norm_Basis[OF k]
-              apply auto
-              done
-            also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
-              apply (rule add_strict_left_mono)
-              using as e
-              apply (auto simp add: field_simps dist_norm)
-              done
-            finally show "y \<in> ball x e"
-              unfolding mem_ball dist_norm using e by (auto simp add:field_simps)
-          qed
-          ultimately show ?thesis
-            apply (rule_tac x="?z" in exI)
-            unfolding Union_insert
-            apply auto
-            done
-        next
-          let ?z = "x + (e/2) *\<^sub>R k"
-          assume as: "x\<bullet>k = b\<bullet>k"
-          have "ball ?z (e / 2) \<inter> i = {}"
-          proof (clarsimp simp only: all_not_in_conv [symmetric])
-            fix y
-            assume "y \<in> ball ?z (e / 2)" and yi: "y \<in> i"
-            then have "dist ?z y < e/2"
-              by auto
-            then have "\<bar>(?z - y) \<bullet> k\<bar> < e/2"
-              using Basis_le_norm[OF k, of "?z - y"]
-              unfolding dist_norm by auto
-            then have "y\<bullet>k > b\<bullet>k"
-              using e k
-              by (auto simp add:field_simps inner_simps inner_Basis as)
-            then have "y \<notin> i"
-              unfolding ab mem_box by (auto intro!: bexI[OF _ k])
-            then show False using yi by auto
-          qed
-          moreover
-          have "ball ?z (e/2) \<subseteq> s \<inter> (\<Union>insert i f)"
-            apply (rule order_trans[OF _ e[THEN conjunct2, unfolded lem1]])
-          proof
-            fix y
-            assume as: "y\<in> ball ?z (e/2)"
-            have "norm (x - y) \<le> \<bar>e\<bar> / 2 + norm (x - y + (e / 2) *\<^sub>R k)"
-              apply (rule order_trans,rule norm_triangle_sub[of "x - y" "- (e/2) *\<^sub>R k"])
-              unfolding norm_scaleR
-              apply (auto simp: k)
-              done
-            also have "\<dots> < \<bar>e\<bar> / 2 + \<bar>e\<bar> / 2"
-              apply (rule add_strict_left_mono)
-              using as unfolding mem_ball dist_norm
-              using e apply (auto simp add: field_simps)
-              done
-            finally show "y \<in> ball x e"
-              unfolding mem_ball dist_norm using e by (auto simp add:field_simps)
-          qed
-          ultimately show ?thesis
-            apply (rule_tac x="?z" in exI)
-            unfolding Union_insert
-            apply auto
-            done
-        qed
-        then obtain x where "ball x (e / 2) \<subseteq> s \<inter> \<Union>f" ..
-        then have "x \<in> s \<inter> interior (\<Union>f)"
-          unfolding lem1[where U="\<Union>f", symmetric]
-          using centre_in_ball e by auto
-        then show ?thesis
-          using insert.hyps(3) insert.prems(1) by blast
-      qed
-    qed
-  qed
-  from this[OF assms(1,3)] x
-  obtain t x e where "t \<in> f" "0 < e" "ball x e \<subseteq> s \<inter> t"
-    by blast
-  then have "x \<in> s" "x \<in> interior t"
-    using open_subset_interior[OF open_ball, of x e t]
-    by auto
-  then show False
-    using \<open>t \<in> f\<close> assms(4) by auto
-qed
-
-subsection \<open>Bounds on intervals where they exist.\<close>
-
-definition interval_upperbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
-  where "interval_upperbound s = (\<Sum>i\<in>Basis. (SUP x:s. x\<bullet>i) *\<^sub>R i)"
-
-definition interval_lowerbound :: "('a::euclidean_space) set \<Rightarrow> 'a"
-   where "interval_lowerbound s = (\<Sum>i\<in>Basis. (INF x:s. x\<bullet>i) *\<^sub>R i)"
-
-lemma interval_upperbound[simp]:
-  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
-    interval_upperbound (cbox a b) = (b::'a::euclidean_space)"
-  unfolding interval_upperbound_def euclidean_representation_setsum cbox_def
-  by (safe intro!: cSup_eq) auto
-
-lemma interval_lowerbound[simp]:
-  "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow>
-    interval_lowerbound (cbox a b) = (a::'a::euclidean_space)"
-  unfolding interval_lowerbound_def euclidean_representation_setsum cbox_def
-  by (safe intro!: cInf_eq) auto
-
-lemmas interval_bounds = interval_upperbound interval_lowerbound
-
-lemma
-  fixes X::"real set"
-  shows interval_upperbound_real[simp]: "interval_upperbound X = Sup X"
-    and interval_lowerbound_real[simp]: "interval_lowerbound X = Inf X"
-  by (auto simp: interval_upperbound_def interval_lowerbound_def)
-
-lemma interval_bounds'[simp]:
-  assumes "cbox a b \<noteq> {}"
-  shows "interval_upperbound (cbox a b) = b"
-    and "interval_lowerbound (cbox a b) = a"
-  using assms unfolding box_ne_empty by auto
-
-
-lemma interval_upperbound_Times:
-  assumes "A \<noteq> {}" and "B \<noteq> {}"
-  shows "interval_upperbound (A \<times> B) = (interval_upperbound A, interval_upperbound B)"
-proof-
-  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:A. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
-  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (SUP x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (SUP x:B. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
-  ultimately show ?thesis unfolding interval_upperbound_def
-      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
-qed
-
-lemma interval_lowerbound_Times:
-  assumes "A \<noteq> {}" and "B \<noteq> {}"
-  shows "interval_lowerbound (A \<times> B) = (interval_lowerbound A, interval_lowerbound B)"
-proof-
-  from assms have fst_image_times': "A = fst ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (i, 0)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:A. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) fst_image_times') (simp del: fst_image_times add: o_def inner_Pair_0)
-  moreover from assms have snd_image_times': "B = snd ` (A \<times> B)" by simp
-  have "(\<Sum>i\<in>Basis. (INF x:A \<times> B. x \<bullet> (0, i)) *\<^sub>R i) = (\<Sum>i\<in>Basis. (INF x:B. x \<bullet> i) *\<^sub>R i)"
-      by (subst (2) snd_image_times') (simp del: snd_image_times add: o_def inner_Pair_0)
-  ultimately show ?thesis unfolding interval_lowerbound_def
-      by (subst setsum_Basis_prod_eq) (auto simp add: setsum_prod)
-qed
-
-subsection \<open>Content (length, area, volume...) of an interval.\<close>
-
-definition "content (s::('a::euclidean_space) set) =
-  (if s = {} then 0 else (\<Prod>i\<in>Basis. (interval_upperbound s)\<bullet>i - (interval_lowerbound s)\<bullet>i))"
-
-lemma interval_not_empty: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> cbox a b \<noteq> {}"
-  unfolding box_eq_empty unfolding not_ex not_less by auto
-
-lemma content_cbox:
-  fixes a :: "'a::euclidean_space"
-  assumes "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-  shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
-  using interval_not_empty[OF assms]
-  unfolding content_def
-  by auto
-
-lemma content_cbox':
-  fixes a :: "'a::euclidean_space"
-  assumes "cbox a b \<noteq> {}"
-  shows "content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
-    using assms box_ne_empty(1) content_cbox by blast
-
-lemma content_real: "a \<le> b \<Longrightarrow> content {a..b} = b - a"
-  by (auto simp: interval_upperbound_def interval_lowerbound_def content_def)
-
-lemma abs_eq_content: "\<bar>y - x\<bar> = (if x\<le>y then content {x .. y} else content {y..x})"
-  by (auto simp: content_real)
-
-lemma content_singleton[simp]: "content {a} = 0"
-proof -
-  have "content (cbox a a) = 0"
-    by (subst content_cbox) (auto simp: ex_in_conv)
-  then show ?thesis by (simp add: cbox_sing)
-qed
-
-lemma content_unit[iff]: "content(cbox 0 (One::'a::euclidean_space)) = 1"
- proof -
-   have *: "\<forall>i\<in>Basis. (0::'a)\<bullet>i \<le> (One::'a)\<bullet>i"
-    by auto
-  have "0 \<in> cbox 0 (One::'a)"
-    unfolding mem_box by auto
-  then show ?thesis
-     unfolding content_def interval_bounds[OF *] using setprod.neutral_const by auto
- qed
-
-lemma content_pos_le[intro]:
-  fixes a::"'a::euclidean_space"
-  shows "0 \<le> content (cbox a b)"
-proof (cases "cbox a b = {}")
-  case False
-  then have *: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-    unfolding box_ne_empty .
-  have "0 \<le> (\<Prod>i\<in>Basis. interval_upperbound (cbox a b) \<bullet> i - interval_lowerbound (cbox a b) \<bullet> i)"
-    apply (rule setprod_nonneg)
-    unfolding interval_bounds[OF *]
-    using *
-    apply auto
-    done
-  also have "\<dots> = content (cbox a b)" using False by (simp add: content_def)
-  finally show ?thesis .
-qed (simp add: content_def)
-
-corollary content_nonneg [simp]:
-  fixes a::"'a::euclidean_space"
-  shows "~ content (cbox a b) < 0"
-using not_le by blast
-
-lemma content_pos_lt:
-  fixes a :: "'a::euclidean_space"
-  assumes "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
-  shows "0 < content (cbox a b)"
-  using assms
-  by (auto simp: content_def box_eq_empty intro!: setprod_pos)
-
-lemma content_eq_0:
-  "content (cbox a b) = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)"
-  by (auto simp: content_def box_eq_empty intro!: setprod_pos bexI)
-
-lemma cond_cases: "(P \<Longrightarrow> Q x) \<Longrightarrow> (\<not> P \<Longrightarrow> Q y) \<Longrightarrow> Q (if P then x else y)"
-  by auto
-
-lemma content_cbox_cases:
-  "content (cbox a (b::'a::euclidean_space)) =
-    (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis else 0)"
-  by (auto simp: not_le content_eq_0 intro: less_imp_le content_cbox)
-
-lemma content_eq_0_interior: "content (cbox a b) = 0 \<longleftrightarrow> interior(cbox a b) = {}"
-  unfolding content_eq_0 interior_cbox box_eq_empty
-  by auto
-
-lemma content_pos_lt_eq:
-  "0 < content (cbox a (b::'a::euclidean_space)) \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)"
-proof (rule iffI)
-  assume "0 < content (cbox a b)"
-  then have "content (cbox a b) \<noteq> 0" by auto
-  then show "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"
-    unfolding content_eq_0 not_ex not_le by fastforce
-next
-  assume "\<forall>i\<in>Basis. a \<bullet> i < b \<bullet> i"
-  then show "0 < content (cbox a b)"
-    by (metis content_pos_lt)
-qed
-
-lemma content_empty [simp]: "content {} = 0"
-  unfolding content_def by auto
-
-lemma content_real_if [simp]: "content {a..b} = (if a \<le> b then b - a else 0)"
-  by (simp add: content_real)
-
-lemma content_subset:
-  assumes "cbox a b \<subseteq> cbox c d"
-  shows "content (cbox a b) \<le> content (cbox c d)"
-proof (cases "cbox a b = {}")
-  case True
-  then show ?thesis
-    using content_pos_le[of c d] by auto
-next
-  case False
-  then have ab_ne: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-    unfolding box_ne_empty by auto
-  then have ab_ab: "a\<in>cbox a b" "b\<in>cbox a b"
-    unfolding mem_box by auto
-  have "cbox c d \<noteq> {}" using assms False by auto
-  then have cd_ne: "\<forall>i\<in>Basis. c \<bullet> i \<le> d \<bullet> i"
-    using assms unfolding box_ne_empty by auto
-  have "\<And>i. i \<in> Basis \<Longrightarrow> 0 \<le> b \<bullet> i - a \<bullet> i"
-    using ab_ne by auto
-  moreover
-  have "\<And>i. i \<in> Basis \<Longrightarrow> b \<bullet> i - a \<bullet> i \<le> d \<bullet> i - c \<bullet> i"
-    using assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(2)]
-          assms[unfolded subset_eq mem_box,rule_format,OF ab_ab(1)]
-      by (metis diff_mono)
-  ultimately show ?thesis
-    unfolding content_def interval_bounds[OF ab_ne] interval_bounds[OF cd_ne]
-    by (simp add: setprod_mono if_not_P[OF False] if_not_P[OF \<open>cbox c d \<noteq> {}\<close>])
-qed
-
-lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0"
-  unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce
-
-lemma content_times[simp]: "content (A \<times> B) = content A * content B"
-proof (cases "A \<times> B = {}")
-  let ?ub1 = "interval_upperbound" and ?lb1 = "interval_lowerbound"
-  let ?ub2 = "interval_upperbound" and ?lb2 = "interval_lowerbound"
-  assume nonempty: "A \<times> B \<noteq> {}"
-  hence "content (A \<times> B) = (\<Prod>i\<in>Basis. (?ub1 A, ?ub2 B) \<bullet> i - (?lb1 A, ?lb2 B) \<bullet> i)"
-      unfolding content_def by (simp add: interval_upperbound_Times interval_lowerbound_Times)
-  also have "... = content A * content B" unfolding content_def using nonempty
-    apply (subst Basis_prod_def, subst setprod.union_disjoint, force, force, force, simp)
-    apply (subst (1 2) setprod.reindex, auto intro: inj_onI)
-    done
-  finally show ?thesis .
-qed (auto simp: content_def)
-
-lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)"
-  by (simp add: cbox_Pair_eq)
-
-lemma content_cbox_pair_eq0_D:
-   "content (cbox (a,c) (b,d)) = 0 \<Longrightarrow> content (cbox a b) = 0 \<or> content (cbox c d) = 0"
-  by (simp add: content_Pair)
-
-lemma content_eq_0_gen:
-  fixes s :: "'a::euclidean_space set"
-  assumes "bounded s"
-  shows "content s = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. \<exists>v. \<forall>x \<in> s. x \<bullet> i = v)"  (is "_ = ?rhs")
-proof safe
-  assume "content s = 0" then show ?rhs
-    apply (clarsimp simp: ex_in_conv content_def split: if_split_asm)
-    apply (rule_tac x=a in bexI)
-    apply (rule_tac x="interval_lowerbound s \<bullet> a" in exI)
-    apply (clarsimp simp: interval_upperbound_def interval_lowerbound_def)
-    apply (drule cSUP_eq_cINF_D)
-    apply (auto simp: bounded_inner_imp_bdd_above [OF assms]  bounded_inner_imp_bdd_below [OF assms])
-    done
-next
-  fix i a
-  assume "i \<in> Basis" "\<forall>x\<in>s. x \<bullet> i = a"
-  then show "content s = 0"
-    apply (clarsimp simp: content_def)
-    apply (rule_tac x=i in bexI)
-    apply (auto simp: interval_upperbound_def interval_lowerbound_def)
-    done
-qed
-
-lemma content_0_subset_gen:
-  fixes a :: "'a::euclidean_space"
-  assumes "content t = 0" "s \<subseteq> t" "bounded t" shows "content s = 0"
-proof -
-  have "bounded s"
-    using assms by (metis bounded_subset)
-  then show ?thesis
-    using assms
-    by (auto simp: content_eq_0_gen)
-qed
-
-lemma content_0_subset: "\<lbrakk>content(cbox a b) = 0; s \<subseteq> cbox a b\<rbrakk> \<Longrightarrow> content s = 0"
-  by (simp add: content_0_subset_gen bounded_cbox)
-
-
-lemma interval_split:
-  fixes a :: "'a::euclidean_space"
-  assumes "k \<in> Basis"
-  shows
-    "cbox a b \<inter> {x. x\<bullet>k \<le> c} = cbox a (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i) *\<^sub>R i)"
-    "cbox a b \<inter> {x. x\<bullet>k \<ge> c} = cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i) *\<^sub>R i) b"
-  apply (rule_tac[!] set_eqI)
-  unfolding Int_iff mem_box mem_Collect_eq
-  using assms
-  apply auto
-  done
-
-lemma content_split:
-  fixes a :: "'a::euclidean_space"
-  assumes "k \<in> Basis"
-  shows "content (cbox a b) = content(cbox a b \<inter> {x. x\<bullet>k \<le> c}) + content(cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-proof cases
-  note simps = interval_split[OF assms] content_cbox_cases
-  have *: "Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}"
-    using assms by auto
-  have *: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
-    "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
-    apply (subst *(1))
-    defer
-    apply (subst *(1))
-    unfolding setprod.insert[OF *(2-)]
-    apply auto
-    done
-  assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-  moreover
-  have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c \<Longrightarrow>
-    x * (b\<bullet>k - a\<bullet>k) = x * (max (a \<bullet> k) c - a \<bullet> k) + x * (b \<bullet> k - max (a \<bullet> k) c)"
-    by  (auto simp add: field_simps)
-  moreover
-  have **: "(\<Prod>i\<in>Basis. ((\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i - a \<bullet> i)) =
-      (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)"
-    "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) =
-      (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))"
-    by (auto intro!: setprod.cong)
-  have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False"
-    unfolding not_le
-    using as[unfolded ,rule_format,of k] assms
-    by auto
-  ultimately show ?thesis
-    using assms
-    unfolding simps **
-    unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"]
-    unfolding *(2)
-    by auto
-next
-  assume "\<not> (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)"
-  then have "cbox a b = {}"
-    unfolding box_eq_empty by (auto simp: not_le)
-  then show ?thesis
-    by (auto simp: not_le)
-qed
-
-subsection \<open>The notion of a gauge --- simply an open set containing the point.\<close>
-
-definition "gauge d \<longleftrightarrow> (\<forall>x. x \<in> d x \<and> open (d x))"
-
-lemma gaugeI:
-  assumes "\<And>x. x \<in> g x"
-    and "\<And>x. open (g x)"
-  shows "gauge g"
-  using assms unfolding gauge_def by auto
-
-lemma gaugeD[dest]:
-  assumes "gauge d"
-  shows "x \<in> d x"
-    and "open (d x)"
-  using assms unfolding gauge_def by auto
-
-lemma gauge_ball_dependent: "\<forall>x. 0 < e x \<Longrightarrow> gauge (\<lambda>x. ball x (e x))"
-  unfolding gauge_def by auto
-
-lemma gauge_ball[intro]: "0 < e \<Longrightarrow> gauge (\<lambda>x. ball x e)"
-  unfolding gauge_def by auto
-
-lemma gauge_trivial[intro!]: "gauge (\<lambda>x. ball x 1)"
-  by (rule gauge_ball) auto
-
-lemma gauge_inter[intro]: "gauge d1 \<Longrightarrow> gauge d2 \<Longrightarrow> gauge (\<lambda>x. d1 x \<inter> d2 x)"
-  unfolding gauge_def by auto
-
-lemma gauge_inters:
-  assumes "finite s"
-    and "\<forall>d\<in>s. gauge (f d)"
-  shows "gauge (\<lambda>x. \<Inter>{f d x | d. d \<in> s})"
-proof -
-  have *: "\<And>x. {f d x |d. d \<in> s} = (\<lambda>d. f d x) ` s"
-    by auto
-  show ?thesis
-    unfolding gauge_def unfolding *
-    using assms unfolding Ball_def Inter_iff mem_Collect_eq gauge_def by auto
-qed
-
-lemma gauge_existence_lemma:
-  "(\<forall>x. \<exists>d :: real. p x \<longrightarrow> 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. p x \<longrightarrow> q d x)"
-  by (metis zero_less_one)
-
-
-subsection \<open>Divisions.\<close>
-
-definition division_of (infixl "division'_of" 40)
-where
-  "s division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>k\<in>s. k \<subseteq> i \<and> k \<noteq> {} \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>k1\<in>s. \<forall>k2\<in>s. k1 \<noteq> k2 \<longrightarrow> interior(k1) \<inter> interior(k2) = {}) \<and>
-    (\<Union>s = i)"
-
-lemma division_ofD[dest]:
-  assumes "s division_of i"
-  shows "finite s"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
-    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior(k1) \<inter> interior(k2) = {}"
-    and "\<Union>s = i"
-  using assms unfolding division_of_def by auto
-
-lemma division_ofI:
-  assumes "finite s"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>k. k \<in> s \<Longrightarrow> k \<noteq> {}"
-    and "\<And>k. k \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>k1 k2. k1 \<in> s \<Longrightarrow> k2 \<in> s \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-    and "\<Union>s = i"
-  shows "s division_of i"
-  using assms unfolding division_of_def by auto
-
-lemma division_of_finite: "s division_of i \<Longrightarrow> finite s"
-  unfolding division_of_def by auto
-
-lemma division_of_self[intro]: "cbox a b \<noteq> {} \<Longrightarrow> {cbox a b} division_of (cbox a b)"
-  unfolding division_of_def by auto
-
-lemma division_of_trivial[simp]: "s division_of {} \<longleftrightarrow> s = {}"
-  unfolding division_of_def by auto
-
-lemma division_of_sing[simp]:
-  "s division_of cbox a (a::'a::euclidean_space) \<longleftrightarrow> s = {cbox a a}"
-  (is "?l = ?r")
-proof
-  assume ?r
-  moreover
-  { fix k
-    assume "s = {{a}}" "k\<in>s"
-    then have "\<exists>x y. k = cbox x y"
-      apply (rule_tac x=a in exI)+
-      apply (force simp: cbox_sing)
-      done
-  }
-  ultimately show ?l
-    unfolding division_of_def cbox_sing by auto
-next
-  assume ?l
-  note * = conjunctD4[OF this[unfolded division_of_def cbox_sing]]
-  {
-    fix x
-    assume x: "x \<in> s" have "x = {a}"
-      using *(2)[rule_format,OF x] by auto
-  }
-  moreover have "s \<noteq> {}"
-    using *(4) by auto
-  ultimately show ?r
-    unfolding cbox_sing by auto
-qed
-
-lemma elementary_empty: obtains p where "p division_of {}"
-  unfolding division_of_trivial by auto
-
-lemma elementary_interval: obtains p where "p division_of (cbox a b)"
-  by (metis division_of_trivial division_of_self)
-
-lemma division_contains: "s division_of i \<Longrightarrow> \<forall>x\<in>i. \<exists>k\<in>s. x \<in> k"
-  unfolding division_of_def by auto
-
-lemma forall_in_division:
-  "d division_of i \<Longrightarrow> (\<forall>x\<in>d. P x) \<longleftrightarrow> (\<forall>a b. cbox a b \<in> d \<longrightarrow> P (cbox a b))"
-  unfolding division_of_def by fastforce
-
-lemma division_of_subset:
-  assumes "p division_of (\<Union>p)"
-    and "q \<subseteq> p"
-  shows "q division_of (\<Union>q)"
-proof (rule division_ofI)
-  note * = division_ofD[OF assms(1)]
-  show "finite q"
-    using "*"(1) assms(2) infinite_super by auto
-  {
-    fix k
-    assume "k \<in> q"
-    then have kp: "k \<in> p"
-      using assms(2) by auto
-    show "k \<subseteq> \<Union>q"
-      using \<open>k \<in> q\<close> by auto
-    show "\<exists>a b. k = cbox a b"
-      using *(4)[OF kp] by auto
-    show "k \<noteq> {}"
-      using *(3)[OF kp] by auto
-  }
-  fix k1 k2
-  assume "k1 \<in> q" "k2 \<in> q" "k1 \<noteq> k2"
-  then have **: "k1 \<in> p" "k2 \<in> p" "k1 \<noteq> k2"
-    using assms(2) by auto
-  show "interior k1 \<inter> interior k2 = {}"
-    using *(5)[OF **] by auto
-qed auto
-
-lemma division_of_union_self[intro]: "p division_of s \<Longrightarrow> p division_of (\<Union>p)"
-  unfolding division_of_def by auto
-
-lemma division_of_content_0:
-  assumes "content (cbox a b) = 0" "d division_of (cbox a b)"
-  shows "\<forall>k\<in>d. content k = 0"
-  unfolding forall_in_division[OF assms(2)]
-  by (metis antisym_conv assms content_pos_le content_subset division_ofD(2))
-
-lemma division_inter:
-  fixes s1 s2 :: "'a::euclidean_space set"
-  assumes "p1 division_of s1"
-    and "p2 division_of s2"
-  shows "{k1 \<inter> k2 | k1 k2 .k1 \<in> p1 \<and> k2 \<in> p2 \<and> k1 \<inter> k2 \<noteq> {}} division_of (s1 \<inter> s2)"
-  (is "?A' division_of _")
-proof -
-  let ?A = "{s. s \<in>  (\<lambda>(k1,k2). k1 \<inter> k2) ` (p1 \<times> p2) \<and> s \<noteq> {}}"
-  have *: "?A' = ?A" by auto
-  show ?thesis
-    unfolding *
-  proof (rule division_ofI)
-    have "?A \<subseteq> (\<lambda>(x, y). x \<inter> y) ` (p1 \<times> p2)"
-      by auto
-    moreover have "finite (p1 \<times> p2)"
-      using assms unfolding division_of_def by auto
-    ultimately show "finite ?A" by auto
-    have *: "\<And>s. \<Union>{x\<in>s. x \<noteq> {}} = \<Union>s"
-      by auto
-    show "\<Union>?A = s1 \<inter> s2"
-      apply (rule set_eqI)
-      unfolding * and UN_iff
-      using division_ofD(6)[OF assms(1)] and division_ofD(6)[OF assms(2)]
-      apply auto
-      done
-    {
-      fix k
-      assume "k \<in> ?A"
-      then obtain k1 k2 where k: "k = k1 \<inter> k2" "k1 \<in> p1" "k2 \<in> p2" "k \<noteq> {}"
-        by auto
-      then show "k \<noteq> {}"
-        by auto
-      show "k \<subseteq> s1 \<inter> s2"
-        using division_ofD(2)[OF assms(1) k(2)] and division_ofD(2)[OF assms(2) k(3)]
-        unfolding k by auto
-      obtain a1 b1 where k1: "k1 = cbox a1 b1"
-        using division_ofD(4)[OF assms(1) k(2)] by blast
-      obtain a2 b2 where k2: "k2 = cbox a2 b2"
-        using division_ofD(4)[OF assms(2) k(3)] by blast
-      show "\<exists>a b. k = cbox a b"
-        unfolding k k1 k2 unfolding inter_interval by auto
-    }
-    fix k1 k2
-    assume "k1 \<in> ?A"
-    then obtain x1 y1 where k1: "k1 = x1 \<inter> y1" "x1 \<in> p1" "y1 \<in> p2" "k1 \<noteq> {}"
-      by auto
-    assume "k2 \<in> ?A"
-    then obtain x2 y2 where k2: "k2 = x2 \<inter> y2" "x2 \<in> p1" "y2 \<in> p2" "k2 \<noteq> {}"
-      by auto
-    assume "k1 \<noteq> k2"
-    then have th: "x1 \<noteq> x2 \<or> y1 \<noteq> y2"
-      unfolding k1 k2 by auto
-    have *: "interior x1 \<inter> interior x2 = {} \<or> interior y1 \<inter> interior y2 = {} \<Longrightarrow>
-      interior (x1 \<inter> y1) \<subseteq> interior x1 \<Longrightarrow> interior (x1 \<inter> y1) \<subseteq> interior y1 \<Longrightarrow>
-      interior (x2 \<inter> y2) \<subseteq> interior x2 \<Longrightarrow> interior (x2 \<inter> y2) \<subseteq> interior y2 \<Longrightarrow>
-      interior (x1 \<inter> y1) \<inter> interior (x2 \<inter> y2) = {}" by auto
-    show "interior k1 \<inter> interior k2 = {}"
-      unfolding k1 k2
-      apply (rule *)
-      using assms division_ofD(5) k1 k2(2) k2(3) th apply auto
-      done
-  qed
-qed
-
-lemma division_inter_1:
-  assumes "d division_of i"
-    and "cbox a (b::'a::euclidean_space) \<subseteq> i"
-  shows "{cbox a b \<inter> k | k. k \<in> d \<and> cbox a b \<inter> k \<noteq> {}} division_of (cbox a b)"
-proof (cases "cbox a b = {}")
-  case True
-  show ?thesis
-    unfolding True and division_of_trivial by auto
-next
-  case False
-  have *: "cbox a b \<inter> i = cbox a b" using assms(2) by auto
-  show ?thesis
-    using division_inter[OF division_of_self[OF False] assms(1)]
-    unfolding * by auto
-qed
-
-lemma elementary_inter:
-  fixes s t :: "'a::euclidean_space set"
-  assumes "p1 division_of s"
-    and "p2 division_of t"
-  shows "\<exists>p. p division_of (s \<inter> t)"
-using assms division_inter by blast
-
-lemma elementary_inters:
-  assumes "finite f"
-    and "f \<noteq> {}"
-    and "\<forall>s\<in>f. \<exists>p. p division_of (s::('a::euclidean_space) set)"
-  shows "\<exists>p. p division_of (\<Inter>f)"
-  using assms
-proof (induct f rule: finite_induct)
-  case (insert x f)
-  show ?case
-  proof (cases "f = {}")
-    case True
-    then show ?thesis
-      unfolding True using insert by auto
-  next
-    case False
-    obtain p where "p division_of \<Inter>f"
-      using insert(3)[OF False insert(5)[unfolded ball_simps,THEN conjunct2]] ..
-    moreover obtain px where "px division_of x"
-      using insert(5)[rule_format,OF insertI1] ..
-    ultimately show ?thesis
-      by (simp add: elementary_inter Inter_insert)
-  qed
-qed auto
-
-lemma division_disjoint_union:
-  assumes "p1 division_of s1"
-    and "p2 division_of s2"
-    and "interior s1 \<inter> interior s2 = {}"
-  shows "(p1 \<union> p2) division_of (s1 \<union> s2)"
-proof (rule division_ofI)
-  note d1 = division_ofD[OF assms(1)]
-  note d2 = division_ofD[OF assms(2)]
-  show "finite (p1 \<union> p2)"
-    using d1(1) d2(1) by auto
-  show "\<Union>(p1 \<union> p2) = s1 \<union> s2"
-    using d1(6) d2(6) by auto
-  {
-    fix k1 k2
-    assume as: "k1 \<in> p1 \<union> p2" "k2 \<in> p1 \<union> p2" "k1 \<noteq> k2"
-    moreover
-    let ?g="interior k1 \<inter> interior k2 = {}"
-    {
-      assume as: "k1\<in>p1" "k2\<in>p2"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(1)]] interior_mono[OF d2(2)[OF as(2)]]
-        using assms(3) by blast
-    }
-    moreover
-    {
-      assume as: "k1\<in>p2" "k2\<in>p1"
-      have ?g
-        using interior_mono[OF d1(2)[OF as(2)]] interior_mono[OF d2(2)[OF as(1)]]
-        using assms(3) by blast
-    }
-    ultimately show ?g
-      using d1(5)[OF _ _ as(3)] and d2(5)[OF _ _ as(3)] by auto
-  }
-  fix k
-  assume k: "k \<in> p1 \<union> p2"
-  show "k \<subseteq> s1 \<union> s2"
-    using k d1(2) d2(2) by auto
-  show "k \<noteq> {}"
-    using k d1(3) d2(3) by auto
-  show "\<exists>a b. k = cbox a b"
-    using k d1(4) d2(4) by auto
-qed
-
-lemma partial_division_extend_1:
-  fixes a b c d :: "'a::euclidean_space"
-  assumes incl: "cbox c d \<subseteq> cbox a b"
-    and nonempty: "cbox c d \<noteq> {}"
-  obtains p where "p division_of (cbox a b)" "cbox c d \<in> p"
-proof
-  let ?B = "\<lambda>f::'a\<Rightarrow>'a \<times> 'a.
-    cbox (\<Sum>i\<in>Basis. (fst (f i) \<bullet> i) *\<^sub>R i) (\<Sum>i\<in>Basis. (snd (f i) \<bullet> i) *\<^sub>R i)"
-  define p where "p = ?B ` (Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)})"
-
-  show "cbox c d \<in> p"
-    unfolding p_def
-    by (auto simp add: box_eq_empty cbox_def intro!: image_eqI[where x="\<lambda>(i::'a)\<in>Basis. (c, d)"])
-  {
-    fix i :: 'a
-    assume "i \<in> Basis"
-    with incl nonempty have "a \<bullet> i \<le> c \<bullet> i" "c \<bullet> i \<le> d \<bullet> i" "d \<bullet> i \<le> b \<bullet> i"
-      unfolding box_eq_empty subset_box by (auto simp: not_le)
-  }
-  note ord = this
-
-  show "p division_of (cbox a b)"
-  proof (rule division_ofI)
-    show "finite p"
-      unfolding p_def by (auto intro!: finite_PiE)
-    {
-      fix k
-      assume "k \<in> p"
-      then obtain f where f: "f \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and k: "k = ?B f"
-        by (auto simp: p_def)
-      then show "\<exists>a b. k = cbox a b"
-        by auto
-      have "k \<subseteq> cbox a b \<and> k \<noteq> {}"
-      proof (simp add: k box_eq_empty subset_box not_less, safe)
-        fix i :: 'a
-        assume i: "i \<in> Basis"
-        with f have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-          by (auto simp: PiE_iff)
-        with i ord[of i]
-        show "a \<bullet> i \<le> fst (f i) \<bullet> i" "snd (f i) \<bullet> i \<le> b \<bullet> i" "fst (f i) \<bullet> i \<le> snd (f i) \<bullet> i"
-          by auto
-      qed
-      then show "k \<noteq> {}" "k \<subseteq> cbox a b"
-        by auto
-      {
-        fix l
-        assume "l \<in> p"
-        then obtain g where g: "g \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}" and l: "l = ?B g"
-          by (auto simp: p_def)
-        assume "l \<noteq> k"
-        have "\<exists>i\<in>Basis. f i \<noteq> g i"
-        proof (rule ccontr)
-          assume "\<not> ?thesis"
-          with f g have "f = g"
-            by (auto simp: PiE_iff extensional_def intro!: ext)
-          with \<open>l \<noteq> k\<close> show False
-            by (simp add: l k)
-        qed
-        then obtain i where *: "i \<in> Basis" "f i \<noteq> g i" ..
-        then have "f i = (a, c) \<or> f i = (c, d) \<or> f i = (d, b)"
-                  "g i = (a, c) \<or> g i = (c, d) \<or> g i = (d, b)"
-          using f g by (auto simp: PiE_iff)
-        with * ord[of i] show "interior l \<inter> interior k = {}"
-          by (auto simp add: l k interior_cbox disjoint_interval intro!: bexI[of _ i])
-      }
-      note \<open>k \<subseteq> cbox a b\<close>
-    }
-    moreover
-    {
-      fix x assume x: "x \<in> cbox a b"
-      have "\<forall>i\<in>Basis. \<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
-      proof
-        fix i :: 'a
-        assume "i \<in> Basis"
-        with x ord[of i]
-        have "(a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> c \<bullet> i) \<or> (c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i) \<or>
-            (d \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> b \<bullet> i)"
-          by (auto simp: cbox_def)
-        then show "\<exists>l. x \<bullet> i \<in> {fst l \<bullet> i .. snd l \<bullet> i} \<and> l \<in> {(a, c), (c, d), (d, b)}"
-          by auto
-      qed
-      then obtain f where
-        f: "\<forall>i\<in>Basis. x \<bullet> i \<in> {fst (f i) \<bullet> i..snd (f i) \<bullet> i} \<and> f i \<in> {(a, c), (c, d), (d, b)}"
-        unfolding bchoice_iff ..
-      moreover from f have "restrict f Basis \<in> Basis \<rightarrow>\<^sub>E {(a, c), (c, d), (d, b)}"
-        by auto
-      moreover from f have "x \<in> ?B (restrict f Basis)"
-        by (auto simp: mem_box)
-      ultimately have "\<exists>k\<in>p. x \<in> k"
-        unfolding p_def by blast
-    }
-    ultimately show "\<Union>p = cbox a b"
-      by auto
-  qed
-qed
-
-lemma partial_division_extend_interval:
-  assumes "p division_of (\<Union>p)" "(\<Union>p) \<subseteq> cbox a b"
-  obtains q where "p \<subseteq> q" "q division_of cbox a (b::'a::euclidean_space)"
-proof (cases "p = {}")
-  case True
-  obtain q where "q division_of (cbox a b)"
-    by (rule elementary_interval)
-  then show ?thesis
-    using True that by blast
-next
-  case False
-  note p = division_ofD[OF assms(1)]
-  have div_cbox: "\<forall>k\<in>p. \<exists>q. q division_of cbox a b \<and> k \<in> q"
-  proof
-    fix k
-    assume kp: "k \<in> p"
-    obtain c d where k: "k = cbox c d"
-      using p(4)[OF kp] by blast
-    have *: "cbox c d \<subseteq> cbox a b" "cbox c d \<noteq> {}"
-      using p(2,3)[OF kp, unfolded k] using assms(2)
-      by (blast intro: order.trans)+
-    obtain q where "q division_of cbox a b" "cbox c d \<in> q"
-      by (rule partial_division_extend_1[OF *])
-    then show "\<exists>q. q division_of cbox a b \<and> k \<in> q"
-      unfolding k by auto
-  qed
-  obtain q where q: "\<And>x. x \<in> p \<Longrightarrow> q x division_of cbox a b" "\<And>x. x \<in> p \<Longrightarrow> x \<in> q x"
-    using bchoice[OF div_cbox] by blast
-  { fix x
-    assume x: "x \<in> p"
-    have "q x division_of \<Union>q x"
-      apply (rule division_ofI)
-      using division_ofD[OF q(1)[OF x]]
-      apply auto
-      done }
-  then have "\<And>x. x \<in> p \<Longrightarrow> \<exists>d. d division_of \<Union>(q x - {x})"
-    by (meson Diff_subset division_of_subset)
-  then have "\<exists>d. d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)"
-    apply -
-    apply (rule elementary_inters [OF finite_imageI[OF p(1)]])
-    apply (auto simp: False elementary_inters [OF finite_imageI[OF p(1)]])
-    done
-  then obtain d where d: "d division_of \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p)" ..
-  have "d \<union> p division_of cbox a b"
-  proof -
-    have te: "\<And>s f t. s \<noteq> {} \<Longrightarrow> \<forall>i\<in>s. f i \<union> i = t \<Longrightarrow> t = \<Inter>(f ` s) \<union> \<Union>s" by auto
-    have cbox_eq: "cbox a b = \<Inter>((\<lambda>i. \<Union>(q i - {i})) ` p) \<union> \<Union>p"
-    proof (rule te[OF False], clarify)
-      fix i
-      assume i: "i \<in> p"
-      show "\<Union>(q i - {i}) \<union> i = cbox a b"
-        using division_ofD(6)[OF q(1)[OF i]] using q(2)[OF i] by auto
-    qed
-    { fix k
-      assume k: "k \<in> p"
-      have *: "\<And>u t s. t \<inter> s = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<inter> t = {}"
-        by auto
-      have "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<inter> interior k = {}"
-      proof (rule *[OF inter_interior_unions_intervals])
-        note qk=division_ofD[OF q(1)[OF k]]
-        show "finite (q k - {k})" "open (interior k)" "\<forall>t\<in>q k - {k}. \<exists>a b. t = cbox a b"
-          using qk by auto
-        show "\<forall>t\<in>q k - {k}. interior k \<inter> interior t = {}"
-          using qk(5) using q(2)[OF k] by auto
-        show "interior (\<Inter>i\<in>p. \<Union>(q i - {i})) \<subseteq> interior (\<Union>(q k - {k}))"
-          apply (rule interior_mono)+
-          using k
-          apply auto
-          done
-      qed } note [simp] = this
-    show "d \<union> p division_of (cbox a b)"
-      unfolding cbox_eq
-      apply (rule division_disjoint_union[OF d assms(1)])
-      apply (rule inter_interior_unions_intervals)
-      apply (rule p open_interior ballI)+
-      apply simp_all
-      done
-  qed
-  then show ?thesis
-    by (meson Un_upper2 that)
-qed
-
-lemma elementary_bounded[dest]:
-  fixes s :: "'a::euclidean_space set"
-  shows "p division_of s \<Longrightarrow> bounded s"
-  unfolding division_of_def by (metis bounded_Union bounded_cbox)
-
-lemma elementary_subset_cbox:
-  "p division_of s \<Longrightarrow> \<exists>a b. s \<subseteq> cbox a (b::'a::euclidean_space)"
-  by (meson elementary_bounded bounded_subset_cbox)
-
-lemma division_union_intervals_exists:
-  fixes a b :: "'a::euclidean_space"
-  assumes "cbox a b \<noteq> {}"
-  obtains p where "(insert (cbox a b) p) division_of (cbox a b \<union> cbox c d)"
-proof (cases "cbox c d = {}")
-  case True
-  show ?thesis
-    apply (rule that[of "{}"])
-    unfolding True
-    using assms
-    apply auto
-    done
-next
-  case False
-  show ?thesis
-  proof (cases "cbox a b \<inter> cbox c d = {}")
-    case True
-    then show ?thesis
-      by (metis that False assms division_disjoint_union division_of_self insert_is_Un interior_Int interior_empty)
-  next
-    case False
-    obtain u v where uv: "cbox a b \<inter> cbox c d = cbox u v"
-      unfolding inter_interval by auto
-    have uv_sub: "cbox u v \<subseteq> cbox c d" using uv by auto
-    obtain p where "p division_of cbox c d" "cbox u v \<in> p"
-      by (rule partial_division_extend_1[OF uv_sub False[unfolded uv]])
-    note p = this division_ofD[OF this(1)]
-    have "interior (cbox a b \<inter> \<Union>(p - {cbox u v})) = interior(cbox u v \<inter> \<Union>(p - {cbox u v}))"
-      apply (rule arg_cong[of _ _ interior])
-      using p(8) uv by auto
-    also have "\<dots> = {}"
-      unfolding interior_Int
-      apply (rule inter_interior_unions_intervals)
-      using p(6) p(7)[OF p(2)] p(3)
-      apply auto
-      done
-    finally have [simp]: "interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}" by simp
-    have cbe: "cbox a b \<union> cbox c d = cbox a b \<union> \<Union>(p - {cbox u v})"
-      using p(8) unfolding uv[symmetric] by auto
-    have "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-    proof -
-      have "{cbox a b} division_of cbox a b"
-        by (simp add: assms division_of_self)
-      then show "insert (cbox a b) (p - {cbox u v}) division_of cbox a b \<union> \<Union>(p - {cbox u v})"
-        by (metis (no_types) Diff_subset \<open>interior (cbox a b) \<inter> interior (\<Union>(p - {cbox u v})) = {}\<close> division_disjoint_union division_of_subset insert_is_Un p(1) p(8))
-    qed
-    with that[of "p - {cbox u v}"] show ?thesis by (simp add: cbe)
-  qed
-qed
-
-lemma division_of_unions:
-  assumes "finite f"
-    and "\<And>p. p \<in> f \<Longrightarrow> p division_of (\<Union>p)"
-    and "\<And>k1 k2. k1 \<in> \<Union>f \<Longrightarrow> k2 \<in> \<Union>f \<Longrightarrow> k1 \<noteq> k2 \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-  shows "\<Union>f division_of \<Union>\<Union>f"
-  using assms
-  by (auto intro!: division_ofI)
-
-lemma elementary_union_interval:
-  fixes a b :: "'a::euclidean_space"
-  assumes "p division_of \<Union>p"
-  obtains q where "q division_of (cbox a b \<union> \<Union>p)"
-proof -
-  note assm = division_ofD[OF assms]
-  have lem1: "\<And>f s. \<Union>\<Union>(f ` s) = \<Union>((\<lambda>x. \<Union>(f x)) ` s)"
-    by auto
-  have lem2: "\<And>f s. f \<noteq> {} \<Longrightarrow> \<Union>{s \<union> t |t. t \<in> f} = s \<union> \<Union>f"
-    by auto
-  {
-    presume "p = {} \<Longrightarrow> thesis"
-      "cbox a b = {} \<Longrightarrow> thesis"
-      "cbox a b \<noteq> {} \<Longrightarrow> interior (cbox a b) = {} \<Longrightarrow> thesis"
-      "p \<noteq> {} \<Longrightarrow> interior (cbox a b)\<noteq>{} \<Longrightarrow> cbox a b \<noteq> {} \<Longrightarrow> thesis"
-    then show thesis by auto
-  next
-    assume as: "p = {}"
-    obtain p where "p division_of (cbox a b)"
-      by (rule elementary_interval)
-    then show thesis
-      using as that by auto
-  next
-    assume as: "cbox a b = {}"
-    show thesis
-      using as assms that by auto
-  next
-    assume as: "interior (cbox a b) = {}" "cbox a b \<noteq> {}"
-    show thesis
-      apply (rule that[of "insert (cbox a b) p"],rule division_ofI)
-      unfolding finite_insert
-      apply (rule assm(1)) unfolding Union_insert
-      using assm(2-4) as
-      apply -
-      apply (fast dest: assm(5))+
-      done
-  next
-    assume as: "p \<noteq> {}" "interior (cbox a b) \<noteq> {}" "cbox a b \<noteq> {}"
-    have "\<forall>k\<in>p. \<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
-    proof
-      fix k
-      assume kp: "k \<in> p"
-      from assm(4)[OF kp] obtain c d where "k = cbox c d" by blast
-      then show "\<exists>q. (insert (cbox a b) q) division_of (cbox a b \<union> k)"
-        by (meson as(3) division_union_intervals_exists)
-    qed
-    from bchoice[OF this] obtain q where "\<forall>x\<in>p. insert (cbox a b) (q x) division_of (cbox a b) \<union> x" ..
-    note q = division_ofD[OF this[rule_format]]
-    let ?D = "\<Union>{insert (cbox a b) (q k) | k. k \<in> p}"
-    show thesis
-    proof (rule that[OF division_ofI])
-      have *: "{insert (cbox a b) (q k) |k. k \<in> p} = (\<lambda>k. insert (cbox a b) (q k)) ` p"
-        by auto
-      show "finite ?D"
-        using "*" assm(1) q(1) by auto
-      show "\<Union>?D = cbox a b \<union> \<Union>p"
-        unfolding * lem1
-        unfolding lem2[OF as(1), of "cbox a b", symmetric]
-        using q(6)
-        by auto
-      fix k
-      assume k: "k \<in> ?D"
-      then show "k \<subseteq> cbox a b \<union> \<Union>p"
-        using q(2) by auto
-      show "k \<noteq> {}"
-        using q(3) k by auto
-      show "\<exists>a b. k = cbox a b"
-        using q(4) k by auto
-      fix k'
-      assume k': "k' \<in> ?D" "k \<noteq> k'"
-      obtain x where x: "k \<in> insert (cbox a b) (q x)" "x\<in>p"
-        using k by auto
-      obtain x' where x': "k'\<in>insert (cbox a b) (q x')" "x'\<in>p"
-        using k' by auto
-      show "interior k \<inter> interior k' = {}"
-      proof (cases "x = x'")
-        case True
-        show ?thesis
-          using True k' q(5) x' x by auto
-      next
-        case False
-        {
-          presume "k = cbox a b \<Longrightarrow> ?thesis"
-            and "k' = cbox a b \<Longrightarrow> ?thesis"
-            and "k \<noteq> cbox a b \<Longrightarrow> k' \<noteq> cbox a b \<Longrightarrow> ?thesis"
-          then show ?thesis by linarith
-        next
-          assume as': "k  = cbox a b"
-          show ?thesis
-            using as' k' q(5) x' by blast
-        next
-          assume as': "k' = cbox a b"
-          show ?thesis
-            using as' k'(2) q(5) x by blast
-        }
-        assume as': "k \<noteq> cbox a b" "k' \<noteq> cbox a b"
-        obtain c d where k: "k = cbox c d"
-          using q(4)[OF x(2,1)] by blast
-        have "interior k \<inter> interior (cbox a b) = {}"
-          using as' k'(2) q(5) x by blast
-        then have "interior k \<subseteq> interior x"
-        using interior_subset_union_intervals
-          by (metis as(2) k q(2) x interior_subset_union_intervals)
-        moreover
-        obtain c d where c_d: "k' = cbox c d"
-          using q(4)[OF x'(2,1)] by blast
-        have "interior k' \<inter> interior (cbox a b) = {}"
-          using as'(2) q(5) x' by blast
-        then have "interior k' \<subseteq> interior x'"
-          by (metis as(2) c_d interior_subset_union_intervals q(2) x'(1) x'(2))
-        ultimately show ?thesis
-          using assm(5)[OF x(2) x'(2) False] by auto
-      qed
-    qed
-  }
-qed
-
-lemma elementary_unions_intervals:
-  assumes fin: "finite f"
-    and "\<And>s. s \<in> f \<Longrightarrow> \<exists>a b. s = cbox a (b::'a::euclidean_space)"
-  obtains p where "p division_of (\<Union>f)"
-proof -
-  have "\<exists>p. p division_of (\<Union>f)"
-  proof (induct_tac f rule:finite_subset_induct)
-    show "\<exists>p. p division_of \<Union>{}" using elementary_empty by auto
-  next
-    fix x F
-    assume as: "finite F" "x \<notin> F" "\<exists>p. p division_of \<Union>F" "x\<in>f"
-    from this(3) obtain p where p: "p division_of \<Union>F" ..
-    from assms(2)[OF as(4)] obtain a b where x: "x = cbox a b" by blast
-    have *: "\<Union>F = \<Union>p"
-      using division_ofD[OF p] by auto
-    show "\<exists>p. p division_of \<Union>insert x F"
-      using elementary_union_interval[OF p[unfolded *], of a b]
-      unfolding Union_insert x * by metis
-  qed (insert assms, auto)
-  then show ?thesis
-    using that by auto
-qed
-
-lemma elementary_union:
-  fixes s t :: "'a::euclidean_space set"
-  assumes "ps division_of s" "pt division_of t"
-  obtains p where "p division_of (s \<union> t)"
-proof -
-  have *: "s \<union> t = \<Union>ps \<union> \<Union>pt"
-    using assms unfolding division_of_def by auto
-  show ?thesis
-    apply (rule elementary_unions_intervals[of "ps \<union> pt"])
-    using assms apply auto
-    by (simp add: * that)
-qed
-
-lemma partial_division_extend:
-  fixes t :: "'a::euclidean_space set"
-  assumes "p division_of s"
-    and "q division_of t"
-    and "s \<subseteq> t"
-  obtains r where "p \<subseteq> r" and "r division_of t"
-proof -
-  note divp = division_ofD[OF assms(1)] and divq = division_ofD[OF assms(2)]
-  obtain a b where ab: "t \<subseteq> cbox a b"
-    using elementary_subset_cbox[OF assms(2)] by auto
-  obtain r1 where "p \<subseteq> r1" "r1 division_of (cbox a b)"
-    using assms
-    by (metis ab dual_order.trans partial_division_extend_interval divp(6))
-  note r1 = this division_ofD[OF this(2)]
-  obtain p' where "p' division_of \<Union>(r1 - p)"
-    apply (rule elementary_unions_intervals[of "r1 - p"])
-    using r1(3,6)
-    apply auto
-    done
-  then obtain r2 where r2: "r2 division_of (\<Union>(r1 - p)) \<inter> (\<Union>q)"
-    by (metis assms(2) divq(6) elementary_inter)
-  {
-    fix x
-    assume x: "x \<in> t" "x \<notin> s"
-    then have "x\<in>\<Union>r1"
-      unfolding r1 using ab by auto
-    then obtain r where r: "r \<in> r1" "x \<in> r"
-      unfolding Union_iff ..
-    moreover
-    have "r \<notin> p"
-    proof
-      assume "r \<in> p"
-      then have "x \<in> s" using divp(2) r by auto
-      then show False using x by auto
-    qed
-    ultimately have "x\<in>\<Union>(r1 - p)" by auto
-  }
-  then have *: "t = \<Union>p \<union> (\<Union>(r1 - p) \<inter> \<Union>q)"
-    unfolding divp divq using assms(3) by auto
-  show ?thesis
-    apply (rule that[of "p \<union> r2"])
-    unfolding *
-    defer
-    apply (rule division_disjoint_union)
-    unfolding divp(6)
-    apply(rule assms r2)+
-  proof -
-    have "interior s \<inter> interior (\<Union>(r1-p)) = {}"
-    proof (rule inter_interior_unions_intervals)
-      show "finite (r1 - p)" and "open (interior s)" and "\<forall>t\<in>r1-p. \<exists>a b. t = cbox a b"
-        using r1 by auto
-      have *: "\<And>s. (\<And>x. x \<in> s \<Longrightarrow> False) \<Longrightarrow> s = {}"
-        by auto
-      show "\<forall>t\<in>r1-p. interior s \<inter> interior t = {}"
-      proof
-        fix m x
-        assume as: "m \<in> r1 - p"
-        have "interior m \<inter> interior (\<Union>p) = {}"
-        proof (rule inter_interior_unions_intervals)
-          show "finite p" and "open (interior m)" and "\<forall>t\<in>p. \<exists>a b. t = cbox a b"
-            using divp by auto
-          show "\<forall>t\<in>p. interior m \<inter> interior t = {}"
-            by (metis DiffD1 DiffD2 as r1(1) r1(7) set_rev_mp)
-        qed
-        then show "interior s \<inter> interior m = {}"
-          unfolding divp by auto
-      qed
-    qed
-    then show "interior s \<inter> interior (\<Union>(r1-p) \<inter> (\<Union>q)) = {}"
-      using interior_subset by auto
-  qed auto
-qed
-
-lemma division_split_left_inj:
-  fixes type :: "'a::euclidean_space"
-  assumes "d division_of i"
-    and "k1 \<in> d"
-    and "k2 \<in> d"
-    and "k1 \<noteq> k2"
-    and "k1 \<inter> {x::'a. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
-    and k: "k\<in>Basis"
-  shows "content(k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
-proof -
-  note d=division_ofD[OF assms(1)]
-  have *: "\<And>(a::'a) b c. content (cbox a b \<inter> {x. x\<bullet>k \<le> c}) = 0 \<longleftrightarrow>
-    interior(cbox a b \<inter> {x. x\<bullet>k \<le> c}) = {}"
-    unfolding  interval_split[OF k] content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
-  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
-    by auto
-  show ?thesis
-    unfolding uv1 uv2 *
-    apply (rule **[OF d(5)[OF assms(2-4)]])
-    apply (simp add: uv1)
-    using assms(5) uv1 by auto
-qed
-
-lemma division_split_right_inj:
-  fixes type :: "'a::euclidean_space"
-  assumes "d division_of i"
-    and "k1 \<in> d"
-    and "k2 \<in> d"
-    and "k1 \<noteq> k2"
-    and "k1 \<inter> {x::'a. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
-    and k: "k \<in> Basis"
-  shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
-proof -
-  note d=division_ofD[OF assms(1)]
-  have *: "\<And>a b::'a. \<And>c. content(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = 0 \<longleftrightarrow>
-    interior(cbox a b \<inter> {x. x\<bullet>k \<ge> c}) = {}"
-    unfolding interval_split[OF k] content_eq_0_interior by auto
-  guess u1 v1 using d(4)[OF assms(2)] by (elim exE) note uv1=this
-  guess u2 v2 using d(4)[OF assms(3)] by (elim exE) note uv2=this
-  have **: "\<And>s t u. s \<inter> t = {} \<Longrightarrow> u \<subseteq> s \<Longrightarrow> u \<subseteq> t \<Longrightarrow> u = {}"
-    by auto
-  show ?thesis
-    unfolding uv1 uv2 *
-    apply (rule **[OF d(5)[OF assms(2-4)]])
-    apply (simp add: uv1)
-    using assms(5) uv1 by auto
-qed
-
-
-lemma division_split:
-  fixes a :: "'a::euclidean_space"
-  assumes "p division_of (cbox a b)"
-    and k: "k\<in>Basis"
-  shows "{l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} division_of(cbox a b \<inter> {x. x\<bullet>k \<le> c})"
-      (is "?p1 division_of ?I1")
-    and "{l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> p \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-      (is "?p2 division_of ?I2")
-proof (rule_tac[!] division_ofI)
-  note p = division_ofD[OF assms(1)]
-  show "finite ?p1" "finite ?p2"
-    using p(1) by auto
-  show "\<Union>?p1 = ?I1" "\<Union>?p2 = ?I2"
-    unfolding p(6)[symmetric] by auto
-  {
-    fix k
-    assume "k \<in> ?p1"
-    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
-    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
-    show "k \<subseteq> ?I1"
-      using l p(2) uv by force
-    show  "k \<noteq> {}"
-      by (simp add: l)
-    show  "\<exists>a b. k = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
-    fix k'
-    assume "k' \<in> ?p1"
-    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
-    assume "k \<noteq> k'"
-    then show "interior k \<inter> interior k' = {}"
-      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
-  }
-  {
-    fix k
-    assume "k \<in> ?p2"
-    then guess l unfolding mem_Collect_eq by (elim exE conjE) note l=this
-    guess u v using p(4)[OF l(2)] by (elim exE) note uv=this
-    show "k \<subseteq> ?I2"
-      using l p(2) uv by force
-    show  "k \<noteq> {}"
-      by (simp add: l)
-    show  "\<exists>a b. k = cbox a b"
-      apply (simp add: l uv p(2-3)[OF l(2)])
-      apply (subst interval_split[OF k])
-      apply (auto intro: order.trans)
-      done
-    fix k'
-    assume "k' \<in> ?p2"
-    then guess l' unfolding mem_Collect_eq by (elim exE conjE) note l'=this
-    assume "k \<noteq> k'"
-    then show "interior k \<inter> interior k' = {}"
-      unfolding l l' using p(5)[OF l(2) l'(2)] by auto
-  }
-qed
-
-subsection \<open>Tagged (partial) divisions.\<close>
-
-definition tagged_partial_division_of (infixr "tagged'_partial'_division'_of" 40)
-  where "s tagged_partial_division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
-      interior k1 \<inter> interior k2 = {})"
-
-lemma tagged_partial_division_ofD[dest]:
-  assumes "s tagged_partial_division_of i"
-  shows "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow>
-      (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow> interior k1 \<inter> interior k2 = {}"
-  using assms unfolding tagged_partial_division_of_def by blast+
-
-definition tagged_division_of (infixr "tagged'_division'_of" 40)
-  where "s tagged_division_of i \<longleftrightarrow> s tagged_partial_division_of i \<and> (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-
-lemma tagged_division_of_finite: "s tagged_division_of i \<Longrightarrow> finite s"
-  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
-
-lemma tagged_division_of:
-  "s tagged_division_of i \<longleftrightarrow>
-    finite s \<and>
-    (\<forall>x k. (x, k) \<in> s \<longrightarrow> x \<in> k \<and> k \<subseteq> i \<and> (\<exists>a b. k = cbox a b)) \<and>
-    (\<forall>x1 k1 x2 k2. (x1, k1) \<in> s \<and> (x2, k2) \<in> s \<and> (x1, k1) \<noteq> (x2, k2) \<longrightarrow>
-      interior k1 \<inter> interior k2 = {}) \<and>
-    (\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  unfolding tagged_division_of_def tagged_partial_division_of_def by auto
-
-lemma tagged_division_ofI:
-  assumes "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1,k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
-      interior k1 \<inter> interior k2 = {}"
-    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  shows "s tagged_division_of i"
-  unfolding tagged_division_of
-  using assms
-  apply auto
-  apply fastforce+
-  done
-
-lemma tagged_division_ofD[dest]:  (*FIXME USE A LOCALE*)
-  assumes "s tagged_division_of i"
-  shows "finite s"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> x \<in> k"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> i"
-    and "\<And>x k. (x,k) \<in> s \<Longrightarrow> \<exists>a b. k = cbox a b"
-    and "\<And>x1 k1 x2 k2. (x1, k1) \<in> s \<Longrightarrow> (x2, k2) \<in> s \<Longrightarrow> (x1, k1) \<noteq> (x2, k2) \<Longrightarrow>
-      interior k1 \<inter> interior k2 = {}"
-    and "(\<Union>{k. \<exists>x. (x,k) \<in> s} = i)"
-  using assms unfolding tagged_division_of by blast+
-
-lemma division_of_tagged_division:
-  assumes "s tagged_division_of i"
-  shows "(snd ` s) division_of i"
-proof (rule division_ofI)
-  note assm = tagged_division_ofD[OF assms]
-  show "\<Union>(snd ` s) = i" "finite (snd ` s)"
-    using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<subseteq> i" "k \<noteq> {}" "\<exists>a b. k = cbox a b"
-    using assm by fastforce+
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by blast
-qed
-
-lemma partial_division_of_tagged_division:
-  assumes "s tagged_partial_division_of i"
-  shows "(snd ` s) division_of \<Union>(snd ` s)"
-proof (rule division_ofI)
-  note assm = tagged_partial_division_ofD[OF assms]
-  show "finite (snd ` s)" "\<Union>(snd ` s) = \<Union>(snd ` s)"
-    using assm by auto
-  fix k
-  assume k: "k \<in> snd ` s"
-  then obtain xk where xk: "(xk, k) \<in> s"
-    by auto
-  then show "k \<noteq> {}" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>(snd ` s)"
-    using assm by auto
-  fix k'
-  assume k': "k' \<in> snd ` s" "k \<noteq> k'"
-  from this(1) obtain xk' where xk': "(xk', k') \<in> s"
-    by auto
-  then show "interior k \<inter> interior k' = {}"
-    using assm(5) k'(2) xk by auto
-qed
-
-lemma tagged_partial_division_subset:
-  assumes "s tagged_partial_division_of i"
-    and "t \<subseteq> s"
-  shows "t tagged_partial_division_of i"
-  using assms
-  unfolding tagged_partial_division_of_def
-  using finite_subset[OF assms(2)]
-  by blast
-
-lemma (in comm_monoid_set) over_tagged_division_lemma:
-  assumes "p tagged_division_of i"
-    and "\<And>u v. cbox u v \<noteq> {} \<Longrightarrow> content (cbox u v) = 0 \<Longrightarrow> d (cbox u v) = \<^bold>1"
-  shows "F (\<lambda>(x,k). d k) p = F d (snd ` p)"
-proof -
-  have *: "(\<lambda>(x,k). d k) = d \<circ> snd"
-    unfolding o_def by (rule ext) auto
-  note assm = tagged_division_ofD[OF assms(1)]
-  show ?thesis
-    unfolding *
-  proof (rule reindex_nontrivial[symmetric])
-    show "finite p"
-      using assm by auto
-    fix x y
-    assume "x\<in>p" "y\<in>p" "x\<noteq>y" "snd x = snd y"
-    obtain a b where ab: "snd x = cbox a b"
-      using assm(4)[of "fst x" "snd x"] \<open>x\<in>p\<close> by auto
-    have "(fst x, snd y) \<in> p" "(fst x, snd y) \<noteq> y"
-      by (metis prod.collapse \<open>x\<in>p\<close> \<open>snd x = snd y\<close> \<open>x \<noteq> y\<close>)+
-    with \<open>x\<in>p\<close> \<open>y\<in>p\<close> have "interior (snd x) \<inter> interior (snd y) = {}"
-      by (intro assm(5)[of "fst x" _ "fst y"]) auto
-    then have "content (cbox a b) = 0"
-      unfolding \<open>snd x = snd y\<close>[symmetric] ab content_eq_0_interior by auto
-    then have "d (cbox a b) = \<^bold>1"
-      using assm(2)[of "fst x" "snd x"] \<open>x\<in>p\<close> ab[symmetric] by (intro assms(2)) auto
-    then show "d (snd x) = \<^bold>1"
-      unfolding ab by auto
-  qed
-qed
-
-lemma tag_in_interval: "p tagged_division_of i \<Longrightarrow> (x, k) \<in> p \<Longrightarrow> x \<in> i"
-  by auto
-
-lemma tagged_division_of_empty: "{} tagged_division_of {}"
-  unfolding tagged_division_of by auto
-
-lemma tagged_partial_division_of_trivial[simp]: "p tagged_partial_division_of {} \<longleftrightarrow> p = {}"
-  unfolding tagged_partial_division_of_def by auto
-
-lemma tagged_division_of_trivial[simp]: "p tagged_division_of {} \<longleftrightarrow> p = {}"
-  unfolding tagged_division_of by auto
-
-lemma tagged_division_of_self: "x \<in> cbox a b \<Longrightarrow> {(x,cbox a b)} tagged_division_of (cbox a b)"
-  by (rule tagged_division_ofI) auto
-
-lemma tagged_division_of_self_real: "x \<in> {a .. b::real} \<Longrightarrow> {(x,{a .. b})} tagged_division_of {a .. b}"
-  unfolding box_real[symmetric]
-  by (rule tagged_division_of_self)
-
-lemma tagged_division_union:
-  assumes "p1 tagged_division_of s1"
-    and "p2 tagged_division_of s2"
-    and "interior s1 \<inter> interior s2 = {}"
-  shows "(p1 \<union> p2) tagged_division_of (s1 \<union> s2)"
-proof (rule tagged_division_ofI)
-  note p1 = tagged_division_ofD[OF assms(1)]
-  note p2 = tagged_division_ofD[OF assms(2)]
-  show "finite (p1 \<union> p2)"
-    using p1(1) p2(1) by auto
-  show "\<Union>{k. \<exists>x. (x, k) \<in> p1 \<union> p2} = s1 \<union> s2"
-    using p1(6) p2(6) by blast
-  fix x k
-  assume xk: "(x, k) \<in> p1 \<union> p2"
-  show "x \<in> k" "\<exists>a b. k = cbox a b"
-    using xk p1(2,4) p2(2,4) by auto
-  show "k \<subseteq> s1 \<union> s2"
-    using xk p1(3) p2(3) by blast
-  fix x' k'
-  assume xk': "(x', k') \<in> p1 \<union> p2" "(x, k) \<noteq> (x', k')"
-  have *: "\<And>a b. a \<subseteq> s1 \<Longrightarrow> b \<subseteq> s2 \<Longrightarrow> interior a \<inter> interior b = {}"
-    using assms(3) interior_mono by blast
-  show "interior k \<inter> interior k' = {}"
-    apply (cases "(x, k) \<in> p1")
-    apply (meson "*" UnE assms(1) assms(2) p1(5) tagged_division_ofD(3) xk'(1) xk'(2))
-    by (metis "*" UnE assms(1) assms(2) inf_sup_aci(1) p2(5) tagged_division_ofD(3) xk xk'(1) xk'(2))
-qed
-
-lemma tagged_division_unions:
-  assumes "finite iset"
-    and "\<forall>i\<in>iset. pfn i tagged_division_of i"
-    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior(i1) \<inter> interior(i2) = {}"
-  shows "\<Union>(pfn ` iset) tagged_division_of (\<Union>iset)"
-proof (rule tagged_division_ofI)
-  note assm = tagged_division_ofD[OF assms(2)[rule_format]]
-  show "finite (\<Union>(pfn ` iset))"
-    apply (rule finite_Union)
-    using assms
-    apply auto
-    done
-  have "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>((\<lambda>i. \<Union>{k. \<exists>x. (x, k) \<in> pfn i}) ` iset)"
-    by blast
-  also have "\<dots> = \<Union>iset"
-    using assm(6) by auto
-  finally show "\<Union>{k. \<exists>x. (x, k) \<in> \<Union>(pfn ` iset)} = \<Union>iset" .
-  fix x k
-  assume xk: "(x, k) \<in> \<Union>(pfn ` iset)"
-  then obtain i where i: "i \<in> iset" "(x, k) \<in> pfn i"
-    by auto
-  show "x \<in> k" "\<exists>a b. k = cbox a b" "k \<subseteq> \<Union>iset"
-    using assm(2-4)[OF i] using i(1) by auto
-  fix x' k'
-  assume xk': "(x', k') \<in> \<Union>(pfn ` iset)" "(x, k) \<noteq> (x', k')"
-  then obtain i' where i': "i' \<in> iset" "(x', k') \<in> pfn i'"
-    by auto
-  have *: "\<And>a b. i \<noteq> i' \<Longrightarrow> a \<subseteq> i \<Longrightarrow> b \<subseteq> i' \<Longrightarrow> interior a \<inter> interior b = {}"
-    using i(1) i'(1)
-    using assms(3)[rule_format] interior_mono
-    by blast
-  show "interior k \<inter> interior k' = {}"
-    apply (cases "i = i'")
-    using assm(5) i' i(2) xk'(2) apply blast
-    using "*" assm(3) i' i by auto
-qed
-
-lemma tagged_partial_division_of_union_self:
-  assumes "p tagged_partial_division_of s"
-  shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
-  using tagged_partial_division_ofD[OF assms]
-  apply auto
-  done
-
-lemma tagged_division_of_union_self:
-  assumes "p tagged_division_of s"
-  shows "p tagged_division_of (\<Union>(snd ` p))"
-  apply (rule tagged_division_ofI)
-  using tagged_division_ofD[OF assms]
-  apply auto
-  done
-
-subsection \<open>Functions closed on boxes: morphisms from boxes to monoids\<close>
-
-text \<open>This auxiliary structure is used to sum up over the elements of a division. Main theorem is
-  @{text operative_division}. Instances for the monoid are @{typ "'a option"}, @{typ real}, and
-  @{typ bool}.\<close>
-
-lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
-  using content_empty unfolding empty_as_interval by auto
-
-paragraph \<open>Using additivity of lifted function to encode definedness.\<close>
-
-definition lift_option :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> 'c option"
-where
-  "lift_option f a' b' = Option.bind a' (\<lambda>a. Option.bind b' (\<lambda>b. Some (f a b)))"
-
-lemma lift_option_simps[simp]:
-  "lift_option f (Some a) (Some b) = Some (f a b)"
-  "lift_option f None b' = None"
-  "lift_option f a' None = None"
-  by (auto simp: lift_option_def)
-
-lemma (in comm_monoid) comm_monoid_lift_option: "comm_monoid (lift_option op \<^bold>*) (Some \<^bold>1)"
-  proof qed (auto simp: lift_option_def ac_simps split: bind_split)
-
-lemma (in comm_monoid) comm_monoid_set_lift_option: "comm_monoid_set (lift_option op \<^bold>*) (Some \<^bold>1)"
-  proof qed (auto simp: lift_option_def ac_simps split: bind_split)
-
-lemma comm_monoid_and: "comm_monoid op \<and> True"
-  proof qed auto
-
-lemma comm_monoid_set_and: "comm_monoid_set op \<and> True"
-  proof qed auto
-
-paragraph \<open>Operative\<close>
-
-definition (in comm_monoid) operative :: "('b::euclidean_space set \<Rightarrow> 'a) \<Rightarrow> bool"
-  where "operative g \<longleftrightarrow>
-    (\<forall>a b. content (cbox a b) = 0 \<longrightarrow> g (cbox a b) = \<^bold>1) \<and>
-    (\<forall>a b c. \<forall>k\<in>Basis. g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c}))"
-
-lemma (in comm_monoid) operativeD[dest]:
-  assumes "operative g"
-  shows "\<And>a b. content (cbox a b) = 0 \<Longrightarrow> g (cbox a b) = \<^bold>1"
-    and "\<And>a b c k. k \<in> Basis \<Longrightarrow> g (cbox a b) = g (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<^bold>* g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-  using assms unfolding operative_def by auto
-
-lemma (in comm_monoid) operative_empty: "operative g \<Longrightarrow> g {} = \<^bold>1"
-  unfolding operative_def by (rule property_empty_interval) auto
-
-lemma operative_content[intro]: "add.operative content"
-  by (force simp add: add.operative_def content_split[symmetric])
-
-definition "division_points (k::('a::euclidean_space) set) d =
-   {(j,x). j \<in> Basis \<and> (interval_lowerbound k)\<bullet>j < x \<and> x < (interval_upperbound k)\<bullet>j \<and>
-     (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
-
-lemma division_points_finite:
-  fixes i :: "'a::euclidean_space set"
-  assumes "d division_of i"
-  shows "finite (division_points i d)"
-proof -
-  note assm = division_ofD[OF assms]
-  let ?M = "\<lambda>j. {(j,x)|x. (interval_lowerbound i)\<bullet>j < x \<and> x < (interval_upperbound i)\<bullet>j \<and>
-    (\<exists>i\<in>d. (interval_lowerbound i)\<bullet>j = x \<or> (interval_upperbound i)\<bullet>j = x)}"
-  have *: "division_points i d = \<Union>(?M ` Basis)"
-    unfolding division_points_def by auto
-  show ?thesis
-    unfolding * using assm by auto
-qed
-
-lemma division_points_subset:
-  fixes a :: "'a::euclidean_space"
-  assumes "d division_of (cbox a b)"
-    and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
-    and k: "k \<in> Basis"
-  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l . l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subseteq>
-      division_points (cbox a b) d" (is ?t1)
-    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l . l \<in> d \<and> ~(l \<inter> {x. x\<bullet>k \<ge> c} = {})} \<subseteq>
-      division_points (cbox a b) d" (is ?t2)
-proof -
-  note assm = division_ofD[OF assms(1)]
-  have *: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-    "\<forall>i\<in>Basis. a\<bullet>i \<le> (\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else  b \<bullet> i) *\<^sub>R i) \<bullet> i"
-    "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i \<le> b\<bullet>i"
-    "min (b \<bullet> k) c = c" "max (a \<bullet> k) c = c"
-    using assms using less_imp_le by auto
-  show ?t1 (*FIXME a horrible mess*)
-    unfolding division_points_def interval_split[OF k, of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
-    unfolding *
-    apply (rule subsetI)
-    unfolding mem_Collect_eq split_beta
-    apply (erule bexE conjE)+
-    apply (simp add: )
-    apply (erule exE conjE)+
-  proof
-    fix i l x
-    assume as:
-      "a \<bullet> fst x < snd x" "snd x < (if fst x = k then c else b \<bullet> fst x)"
-      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      "i = l \<inter> {x. x \<bullet> k \<le> c}" "l \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
-      and fstx: "fst x \<in> Basis"
-    from assm(4)[OF this(5)] guess u v apply-by(erule exE)+ note l=this
-    have *: "\<forall>i\<in>Basis. u \<bullet> i \<le> (\<Sum>i\<in>Basis. (if i = k then min (v \<bullet> k) c else v \<bullet> i) *\<^sub>R i) \<bullet> i"
-      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
-      using l using as(6) unfolding box_ne_empty[symmetric] by auto
-    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using as(1-3,5) fstx
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
-      apply (auto split: if_split_asm)
-      done
-    show "snd x < b \<bullet> fst x"
-      using as(2) \<open>c < b\<bullet>k\<close> by (auto split: if_split_asm)
-  qed
-  show ?t2
-    unfolding division_points_def interval_split[OF k, of a b]
-    unfolding interval_bounds[OF *(1)] interval_bounds[OF *(2)] interval_bounds[OF *(3)]
-    unfolding *
-    unfolding subset_eq
-    apply rule
-    unfolding mem_Collect_eq split_beta
-    apply (erule bexE conjE)+
-    apply (simp only: mem_Collect_eq inner_setsum_left_Basis simp_thms)
-    apply (erule exE conjE)+
-  proof
-    fix i l x
-    assume as:
-      "(if fst x = k then c else a \<bullet> fst x) < snd x" "snd x < b \<bullet> fst x"
-      "interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      "i = l \<inter> {x. c \<le> x \<bullet> k}" "l \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}"
-      and fstx: "fst x \<in> Basis"
-    from assm(4)[OF this(5)] guess u v by (elim exE) note l=this
-    have *: "\<forall>i\<in>Basis. (\<Sum>i\<in>Basis. (if i = k then max (u \<bullet> k) c else u \<bullet> i) *\<^sub>R i) \<bullet> i \<le> v \<bullet> i"
-      using as(6) unfolding l interval_split[OF k] box_ne_empty as .
-    have **: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i"
-      using l using as(6) unfolding box_ne_empty[symmetric] by auto
-    show "\<exists>i\<in>d. interval_lowerbound i \<bullet> fst x = snd x \<or> interval_upperbound i \<bullet> fst x = snd x"
-      apply (rule bexI[OF _ \<open>l \<in> d\<close>])
-      using as(1-3,5) fstx
-      unfolding l interval_bounds[OF **] interval_bounds[OF *] interval_split[OF k] as
-      apply (auto split: if_split_asm)
-      done
-    show "a \<bullet> fst x < snd x"
-      using as(1) \<open>a\<bullet>k < c\<close> by (auto split: if_split_asm)
-   qed
-qed
-
-lemma division_points_psubset:
-  fixes a :: "'a::euclidean_space"
-  assumes "d division_of (cbox a b)"
-      and "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i"  "a\<bullet>k < c" "c < b\<bullet>k"
-      and "l \<in> d"
-      and "interval_lowerbound l\<bullet>k = c \<or> interval_upperbound l\<bullet>k = c"
-      and k: "k \<in> Basis"
-  shows "division_points (cbox a b \<inter> {x. x\<bullet>k \<le> c}) {l \<inter> {x. x\<bullet>k \<le> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}} \<subset>
-         division_points (cbox a b) d" (is "?D1 \<subset> ?D")
-    and "division_points (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) {l \<inter> {x. x\<bullet>k \<ge> c} | l. l\<in>d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}} \<subset>
-         division_points (cbox a b) d" (is "?D2 \<subset> ?D")
-proof -
-  have ab: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-    using assms(2) by (auto intro!:less_imp_le)
-  guess u v using division_ofD(4)[OF assms(1,5)] by (elim exE) note l=this
-  have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "\<forall>i\<in>Basis. a\<bullet>i \<le> u\<bullet>i \<and> v\<bullet>i \<le> b\<bullet>i"
-    using division_ofD(2,2,3)[OF assms(1,5)] unfolding l box_ne_empty
-    using subset_box(1)
-    apply auto
-    apply blast+
-    done
-  have *: "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
-          "interval_upperbound (cbox a b \<inter> {x. x \<bullet> k \<le> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
-    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
-    using uv[rule_format, of k] ab k
-    by auto
-  have "\<exists>x. x \<in> ?D - ?D1"
-    using assms(3-)
-    unfolding division_points_def interval_bounds[OF ab]
-    apply -
-    apply (erule disjE)
-    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
-    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
-    done
-  moreover have "?D1 \<subseteq> ?D"
-    by (auto simp add: assms division_points_subset)
-  ultimately show "?D1 \<subset> ?D"
-    by blast
-  have *: "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_lowerbound l \<bullet> k}) \<bullet> k = interval_lowerbound l \<bullet> k"
-    "interval_lowerbound (cbox a b \<inter> {x. x \<bullet> k \<ge> interval_upperbound l \<bullet> k}) \<bullet> k = interval_upperbound l \<bullet> k"
-    unfolding l interval_split[OF k] interval_bounds[OF uv(1)]
-    using uv[rule_format, of k] ab k
-    by auto
-  have "\<exists>x. x \<in> ?D - ?D2"
-    using assms(3-)
-    unfolding division_points_def interval_bounds[OF ab]
-    apply -
-    apply (erule disjE)
-    apply (rule_tac x="(k,(interval_lowerbound l)\<bullet>k)" in exI, force simp add: *)
-    apply (rule_tac x="(k,(interval_upperbound l)\<bullet>k)" in exI, force simp add: *)
-    done
-  moreover have "?D2 \<subseteq> ?D"
-    by (auto simp add: assms division_points_subset)
-  ultimately show "?D2 \<subset> ?D"
-    by blast
-qed
-
-lemma (in comm_monoid_set) operative_division:
-  fixes g :: "'b::euclidean_space set \<Rightarrow> 'a"
-  assumes g: "operative g" and d: "d division_of (cbox a b)" shows "F g d = g (cbox a b)"
-proof -
-  define C where [abs_def]: "C = card (division_points (cbox a b) d)"
-  then show ?thesis
-    using d
-  proof (induction C arbitrary: a b d rule: less_induct)
-    case (less a b d)
-    show ?case
-    proof cases
-      show "content (cbox a b) = 0 \<Longrightarrow> F g d = g (cbox a b)"
-        using division_of_content_0[OF _ less.prems] operativeD(1)[OF  g] division_ofD(4)[OF less.prems]
-        by (fastforce intro!: neutral)
-    next
-      assume "content (cbox a b) \<noteq> 0"
-      note ab = this[unfolded content_lt_nz[symmetric] content_pos_lt_eq]
-      then have ab': "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i"
-        by (auto intro!: less_imp_le)
-      show "F g d = g (cbox a b)"
-      proof (cases "division_points (cbox a b) d = {}")
-        case True
-        { fix u v and j :: 'b
-          assume j: "j \<in> Basis" and as: "cbox u v \<in> d"
-          then have "cbox u v \<noteq> {}"
-            using less.prems by blast
-          then have uv: "\<forall>i\<in>Basis. u\<bullet>i \<le> v\<bullet>i" "u\<bullet>j \<le> v\<bullet>j"
-            using j unfolding box_ne_empty by auto
-          have *: "\<And>p r Q. \<not> j\<in>Basis \<or> p \<or> r \<or> (\<forall>x\<in>d. Q x) \<Longrightarrow> p \<or> r \<or> Q (cbox u v)"
-            using as j by auto
-          have "(j, u\<bullet>j) \<notin> division_points (cbox a b) d"
-               "(j, v\<bullet>j) \<notin> division_points (cbox a b) d" using True by auto
-          note this[unfolded de_Morgan_conj division_points_def mem_Collect_eq split_conv interval_bounds[OF ab'] bex_simps]
-          note *[OF this(1)] *[OF this(2)] note this[unfolded interval_bounds[OF uv(1)]]
-          moreover
-          have "a\<bullet>j \<le> u\<bullet>j" "v\<bullet>j \<le> b\<bullet>j"
-            using division_ofD(2,2,3)[OF \<open>d division_of cbox a b\<close> as]
-            apply (metis j subset_box(1) uv(1))
-            by (metis \<open>cbox u v \<subseteq> cbox a b\<close> j subset_box(1) uv(1))
-          ultimately have "u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j"
-            unfolding not_less de_Morgan_disj using ab[rule_format,of j] uv(2) j by force }
-        then have d': "\<forall>i\<in>d. \<exists>u v. i = cbox u v \<and>
-          (\<forall>j\<in>Basis. u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = a\<bullet>j \<or> u\<bullet>j = b\<bullet>j \<and> v\<bullet>j = b\<bullet>j \<or> u\<bullet>j = a\<bullet>j \<and> v\<bullet>j = b\<bullet>j)"
-          unfolding forall_in_division[OF less.prems] by blast
-        have "(1/2) *\<^sub>R (a+b) \<in> cbox a b"
-          unfolding mem_box using ab by(auto intro!: less_imp_le simp: inner_simps)
-        note this[unfolded division_ofD(6)[OF \<open>d division_of cbox a b\<close>,symmetric] Union_iff]
-        then guess i .. note i=this
-        guess u v using d'[rule_format,OF i(1)] by (elim exE conjE) note uv=this
-        have "cbox a b \<in> d"
-        proof -
-          have "u = a" "v = b"
-            unfolding euclidean_eq_iff[where 'a='b]
-          proof safe
-            fix j :: 'b
-            assume j: "j \<in> Basis"
-            note i(2)[unfolded uv mem_box,rule_format,of j]
-            then show "u \<bullet> j = a \<bullet> j" and "v \<bullet> j = b \<bullet> j"
-              using uv(2)[rule_format,of j] j by (auto simp: inner_simps)
-          qed
-          then have "i = cbox a b" using uv by auto
-          then show ?thesis using i by auto
-        qed
-        then have deq: "d = insert (cbox a b) (d - {cbox a b})"
-          by auto
-        have "F g (d - {cbox a b}) = \<^bold>1"
-        proof (intro neutral ballI)
-          fix x
-          assume x: "x \<in> d - {cbox a b}"
-          then have "x\<in>d"
-            by auto note d'[rule_format,OF this]
-          then guess u v by (elim exE conjE) note uv=this
-          have "u \<noteq> a \<or> v \<noteq> b"
-            using x[unfolded uv] by auto
-          then obtain j where "u\<bullet>j \<noteq> a\<bullet>j \<or> v\<bullet>j \<noteq> b\<bullet>j" and j: "j \<in> Basis"
-            unfolding euclidean_eq_iff[where 'a='b] by auto
-          then have "u\<bullet>j = v\<bullet>j"
-            using uv(2)[rule_format,OF j] by auto
-          then have "content (cbox u v) = 0"
-            unfolding content_eq_0 using j
-            by force
-          then show "g x = \<^bold>1"
-            unfolding uv(1) by (rule operativeD(1)[OF g])
-        qed
-        then show "F g d = g (cbox a b)"
-          using division_ofD[OF less.prems]
-          apply (subst deq)
-          apply (subst insert)
-          apply auto
-          done
-      next
-        case False
-        then have "\<exists>x. x \<in> division_points (cbox a b) d"
-          by auto
-        then guess k c
-          unfolding split_paired_Ex division_points_def mem_Collect_eq split_conv
-          apply (elim exE conjE)
-          done
-        note this(2-4,1) note kc=this[unfolded interval_bounds[OF ab']]
-        from this(3) guess j .. note j=this
-        define d1 where "d1 = {l \<inter> {x. x\<bullet>k \<le> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
-        define d2 where "d2 = {l \<inter> {x. x\<bullet>k \<ge> c} | l. l \<in> d \<and> l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
-        define cb where "cb = (\<Sum>i\<in>Basis. (if i = k then c else b\<bullet>i) *\<^sub>R i)"
-        define ca where "ca = (\<Sum>i\<in>Basis. (if i = k then c else a\<bullet>i) *\<^sub>R i)"
-        note division_points_psubset[OF \<open>d division_of cbox a b\<close> ab kc(1-2) j]
-        note psubset_card_mono[OF _ this(1)] psubset_card_mono[OF _ this(2)]
-        then have *: "F g d1 = g (cbox a b \<inter> {x. x\<bullet>k \<le> c})" "F g d2 = g (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-          unfolding interval_split[OF kc(4)]
-          apply (rule_tac[!] "less.hyps"[rule_format])
-          using division_split[OF \<open>d division_of cbox a b\<close>, where k=k and c=c]
-          apply (simp_all add: interval_split kc d1_def d2_def division_points_finite[OF \<open>d division_of cbox a b\<close>])
-          done
-        { fix l y
-          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. x \<bullet> k \<le> c} = y \<inter> {x. x \<bullet> k \<le> c}" "l \<noteq> y"
-          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
-          have "g (l \<inter> {x. x \<bullet> k \<le> c}) = \<^bold>1"
-            unfolding leq interval_split[OF kc(4)]
-            apply (rule operativeD[OF g])
-            unfolding interval_split[symmetric, OF kc(4)]
-            using division_split_left_inj less as kc leq by blast
-        } note fxk_le = this
-        { fix l y
-          assume as: "l \<in> d" "y \<in> d" "l \<inter> {x. c \<le> x \<bullet> k} = y \<inter> {x. c \<le> x \<bullet> k}" "l \<noteq> y"
-          from division_ofD(4)[OF \<open>d division_of cbox a b\<close> this(1)] guess u v by (elim exE) note leq=this
-          have "g (l \<inter> {x. x \<bullet> k \<ge> c}) = \<^bold>1"
-            unfolding leq interval_split[OF kc(4)]
-            apply (rule operativeD(1)[OF g])
-            unfolding interval_split[symmetric,OF kc(4)]
-            using division_split_right_inj less leq as kc by blast
-        } note fxk_ge = this
-        have d1_alt: "d1 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<le> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
-          using d1_def by auto
-        have d2_alt: "d2 = (\<lambda>l. l \<inter> {x. x\<bullet>k \<ge> c}) ` {l \<in> d. l \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
-          using d2_def by auto
-        have "g (cbox a b) = F g d1 \<^bold>* F g d2" (is "_ = ?prev")
-          unfolding * using g kc(4) by blast
-        also have "F g d1 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d"
-          unfolding d1_alt using division_of_finite[OF less.prems] fxk_le
-          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
-        also have "F g d2 = F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d"
-          unfolding d2_alt using division_of_finite[OF less.prems] fxk_ge
-          by (subst reindex_nontrivial) (auto intro!: mono_neutral_cong_left simp: operative_empty[OF g])
-        also have *: "\<forall>x\<in>d. g x = g (x \<inter> {x. x \<bullet> k \<le> c}) \<^bold>* g (x \<inter> {x. c \<le> x \<bullet> k})"
-          unfolding forall_in_division[OF \<open>d division_of cbox a b\<close>]
-          using g kc(4) by blast
-        have "F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<le> c})) d \<^bold>* F (\<lambda>l. g (l \<inter> {x. x\<bullet>k \<ge> c})) d = F g d"
-          using * by (simp add: distrib)
-        finally show ?thesis by auto
-      qed
-    qed
-  qed
-qed
-
-lemma (in comm_monoid_set) operative_tagged_division:
-  assumes f: "operative g" and d: "d tagged_division_of (cbox a b)"
-  shows "F (\<lambda>(x, l). g l) d = g (cbox a b)"
-  unfolding d[THEN division_of_tagged_division, THEN operative_division[OF f], symmetric]
-  by (simp add: f[THEN operativeD(1)] over_tagged_division_lemma[OF d])
-
-lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> setsum content d = content (cbox a b)"
-  by (metis operative_content setsum.operative_division)
-
-lemma additive_content_tagged_division:
-  "d tagged_division_of (cbox a b) \<Longrightarrow> setsum (\<lambda>(x,l). content l) d = content (cbox a b)"
-  unfolding setsum.operative_tagged_division[OF operative_content, symmetric] by blast
-
-lemma
-  shows real_inner_1_left: "inner 1 x = x"
-  and real_inner_1_right: "inner x 1 = x"
-  by simp_all
-
-lemma content_real_eq_0: "content {a .. b::real} = 0 \<longleftrightarrow> a \<ge> b"
-  by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0)
-
-lemma interval_real_split:
-  "{a .. b::real} \<inter> {x. x \<le> c} = {a .. min b c}"
-  "{a .. b} \<inter> {x. c \<le> x} = {max a c .. b}"
-  apply (metis Int_atLeastAtMostL1 atMost_def)
-  apply (metis Int_atLeastAtMostL2 atLeast_def)
-  done
-
-lemma (in comm_monoid) operative_1_lt:
-  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
-    ((\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1) \<and> (\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
-  apply (simp add: operative_def content_real_eq_0 atMost_def[symmetric] atLeast_def[symmetric]
-              del: content_real_if)
-proof safe
-  fix a b c :: real
-  assume *: "\<forall>a b c. g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
-  assume "a < c" "c < b"
-  with *[rule_format, of a b c] show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    by (simp add: less_imp_le min.absorb2 max.absorb2)
-next
-  fix a b c :: real
-  assume as: "\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1"
-    "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-  from as(1)[rule_format, of 0 1] as(1)[rule_format, of a a for a] as(2)
-  have [simp]: "g {} = \<^bold>1" "\<And>a. g {a} = \<^bold>1"
-    "\<And>a b c. a < c \<Longrightarrow> c < b \<Longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    by auto
-  show "g {a..b} = g {a..min b c} \<^bold>* g {max a c..b}"
-    by (auto simp: min_def max_def le_less)
-qed
-
-lemma (in comm_monoid) operative_1_le:
-  "operative (g :: real set \<Rightarrow> 'a) \<longleftrightarrow>
-    ((\<forall>a b. b \<le> a \<longrightarrow> g {a..b} = \<^bold>1) \<and> (\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a .. c} \<^bold>* g {c .. b} = g {a .. b}))"
-  unfolding operative_1_lt
-proof safe
-  fix a b c :: real
-  assume as: "\<forall>a b c. a \<le> c \<and> c \<le> b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}" "a < c" "c < b"
-  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    apply (rule as(1)[rule_format])
-    using as(2-)
-    apply auto
-    done
-next
-  fix a b c :: real
-  assume "\<forall>a b. b \<le> a \<longrightarrow> g {a .. b} = \<^bold>1"
-    and "\<forall>a b c. a < c \<and> c < b \<longrightarrow> g {a..c} \<^bold>* g {c..b} = g {a..b}"
-    and "a \<le> c"
-    and "c \<le> b"
-  note as = this[rule_format]
-  show "g {a..c} \<^bold>* g {c..b} = g {a..b}"
-  proof (cases "c = a \<or> c = b")
-    case False
-    then show ?thesis
-      apply -
-      apply (subst as(2))
-      using as(3-)
-      apply auto
-      done
-  next
-    case True
-    then show ?thesis
-    proof
-      assume *: "c = a"
-      then have "g {a .. c} = \<^bold>1"
-        apply -
-        apply (rule as(1)[rule_format])
-        apply auto
-        done
-      then show ?thesis
-        unfolding * by auto
-    next
-      assume *: "c = b"
-      then have "g {c .. b} = \<^bold>1"
-        apply -
-        apply (rule as(1)[rule_format])
-        apply auto
-        done
-      then show ?thesis
-        unfolding * by auto
-    qed
-  qed
-qed
-
-subsection \<open>Fine-ness of a partition w.r.t. a gauge.\<close>
-
-definition fine  (infixr "fine" 46)
-  where "d fine s \<longleftrightarrow> (\<forall>(x,k) \<in> s. k \<subseteq> d x)"
-
-lemma fineI:
-  assumes "\<And>x k. (x, k) \<in> s \<Longrightarrow> k \<subseteq> d x"
-  shows "d fine s"
-  using assms unfolding fine_def by auto
-
-lemma fineD[dest]:
-  assumes "d fine s"
-  shows "\<And>x k. (x,k) \<in> s \<Longrightarrow> k \<subseteq> d x"
-  using assms unfolding fine_def by auto
-
-lemma fine_inter: "(\<lambda>x. d1 x \<inter> d2 x) fine p \<longleftrightarrow> d1 fine p \<and> d2 fine p"
-  unfolding fine_def by auto
-
-lemma fine_inters:
- "(\<lambda>x. \<Inter>{f d x | d.  d \<in> s}) fine p \<longleftrightarrow> (\<forall>d\<in>s. (f d) fine p)"
-  unfolding fine_def by blast
-
-lemma fine_union: "d fine p1 \<Longrightarrow> d fine p2 \<Longrightarrow> d fine (p1 \<union> p2)"
-  unfolding fine_def by blast
-
-lemma fine_unions: "(\<And>p. p \<in> ps \<Longrightarrow> d fine p) \<Longrightarrow> d fine (\<Union>ps)"
-  unfolding fine_def by auto
-
-lemma fine_subset: "p \<subseteq> q \<Longrightarrow> d fine q \<Longrightarrow> d fine p"
-  unfolding fine_def by blast
-
-
-subsection \<open>Gauge integral. Define on compact intervals first, then use a limit.\<close>
-
-definition has_integral_compact_interval (infixr "has'_integral'_compact'_interval" 46)
-  where "(f has_integral_compact_interval y) i \<longleftrightarrow>
-    (\<forall>e>0. \<exists>d. gauge d \<and>
-      (\<forall>p. p tagged_division_of i \<and> d fine p \<longrightarrow>
-        norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - y) < e))"
-
-definition has_integral ::
-    "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool"
-  (infixr "has'_integral" 46)
-  where "(f has_integral y) i \<longleftrightarrow>
-    (if \<exists>a b. i = cbox a b
-     then (f has_integral_compact_interval y) i
-     else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral_compact_interval z) (cbox a b) \<and>
-        norm (z - y) < e)))"
-
-lemma has_integral:
-  "(f has_integral y) (cbox a b) \<longleftrightarrow>
-    (\<forall>e>0. \<exists>d. gauge d \<and>
-      (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-        norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
-  unfolding has_integral_def has_integral_compact_interval_def
-  by auto
-
-lemma has_integral_real:
-  "(f has_integral y) {a .. b::real} \<longleftrightarrow>
-    (\<forall>e>0. \<exists>d. gauge d \<and>
-      (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
-        norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f x) p - y) < e))"
-  unfolding box_real[symmetric]
-  by (rule has_integral)
-
-lemma has_integralD[dest]:
-  assumes "(f has_integral y) (cbox a b)"
-    and "e > 0"
-  obtains d where "gauge d"
-    and "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d fine p \<Longrightarrow>
-      norm (setsum (\<lambda>(x,k). content(k) *\<^sub>R f(x)) p - y) < e"
-  using assms unfolding has_integral by auto
-
-lemma has_integral_alt:
-  "(f has_integral y) i \<longleftrightarrow>
-    (if \<exists>a b. i = cbox a b
-     then (f has_integral y) i
-     else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e)))"
-  unfolding has_integral
-  unfolding has_integral_compact_interval_def has_integral_def
-  by auto
-
-lemma has_integral_altD:
-  assumes "(f has_integral y) i"
-    and "\<not> (\<exists>a b. i = cbox a b)"
-    and "e>0"
-  obtains B where "B > 0"
-    and "\<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - y) < e)"
-  using assms
-  unfolding has_integral
-  unfolding has_integral_compact_interval_def has_integral_def
-  by auto
-
-definition integrable_on (infixr "integrable'_on" 46)
-  where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)"
-
-definition "integral i f = (SOME y. (f has_integral y) i \<or> ~ f integrable_on i \<and> y=0)"
-
-lemma integrable_integral[dest]: "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i"
-  unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex)
-
-lemma not_integrable_integral: "~ f integrable_on i \<Longrightarrow> integral i f = 0"
-  unfolding integrable_on_def integral_def by blast
-
-lemma has_integral_integrable[intro]: "(f has_integral i) s \<Longrightarrow> f integrable_on s"
-  unfolding integrable_on_def by auto
-
-lemma has_integral_integral: "f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s"
-  by auto
-
-lemma setsum_content_null:
-  assumes "content (cbox a b) = 0"
-    and "p tagged_division_of (cbox a b)"
-  shows "setsum (\<lambda>(x,k). content k *\<^sub>R f x) p = (0::'a::real_normed_vector)"
-proof (rule setsum.neutral, rule)
-  fix y
-  assume y: "y \<in> p"
-  obtain x k where xk: "y = (x, k)"
-    using surj_pair[of y] by blast
-  note assm = tagged_division_ofD(3-4)[OF assms(2) y[unfolded xk]]
-  from this(2) obtain c d where k: "k = cbox c d" by blast
-  have "(\<lambda>(x, k). content k *\<^sub>R f x) y = content k *\<^sub>R f x"
-    unfolding xk by auto
-  also have "\<dots> = 0"
-    using content_subset[OF assm(1)[unfolded k]] content_pos_le[of c d]
-    unfolding assms(1) k
-    by auto
-  finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" .
-qed
-
-
-subsection \<open>Some basic combining lemmas.\<close>
-
-lemma tagged_division_unions_exists:
-  assumes "finite iset"
-    and "\<forall>i\<in>iset. \<exists>p. p tagged_division_of i \<and> d fine p"
-    and "\<forall>i1\<in>iset. \<forall>i2\<in>iset. i1 \<noteq> i2 \<longrightarrow> interior i1 \<inter> interior i2 = {}"
-    and "\<Union>iset = i"
-   obtains p where "p tagged_division_of i" and "d fine p"
-proof -
-  obtain pfn where pfn:
-    "\<And>x. x \<in> iset \<Longrightarrow> pfn x tagged_division_of x"
-    "\<And>x. x \<in> iset \<Longrightarrow> d fine pfn x"
-    using bchoice[OF assms(2)] by auto
-  show thesis
-    apply (rule_tac p="\<Union>(pfn ` iset)" in that)
-    using assms(1) assms(3) assms(4) pfn(1) tagged_division_unions apply force
-    by (metis (mono_tags, lifting) fine_unions imageE pfn(2))
-qed
-
-
-subsection \<open>The set we're concerned with must be closed.\<close>
-
-lemma division_of_closed:
-  fixes i :: "'n::euclidean_space set"
-  shows "s division_of i \<Longrightarrow> closed i"
-  unfolding division_of_def by fastforce
-
-subsection \<open>General bisection principle for intervals; might be useful elsewhere.\<close>
-
-lemma interval_bisection_step:
-  fixes type :: "'a::euclidean_space"
-  assumes "P {}"
-    and "\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P (s \<union> t)"
-    and "\<not> P (cbox a (b::'a))"
-  obtains c d where "\<not> P (cbox c d)"
-    and "\<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
-proof -
-  have "cbox a b \<noteq> {}"
-    using assms(1,3) by metis
-  then have ab: "\<And>i. i\<in>Basis \<Longrightarrow> a \<bullet> i \<le> b \<bullet> i"
-    by (force simp: mem_box)
-  { fix f
-    have "\<lbrakk>finite f;
-           \<And>s. s\<in>f \<Longrightarrow> P s;
-           \<And>s. s\<in>f \<Longrightarrow> \<exists>a b. s = cbox a b;
-           \<And>s t. s\<in>f \<Longrightarrow> t\<in>f \<Longrightarrow> s \<noteq> t \<Longrightarrow> interior s \<inter> interior t = {}\<rbrakk> \<Longrightarrow> P (\<Union>f)"
-    proof (induct f rule: finite_induct)
-      case empty
-      show ?case
-        using assms(1) by auto
-    next
-      case (insert x f)
-      show ?case
-        unfolding Union_insert
-        apply (rule assms(2)[rule_format])
-        using inter_interior_unions_intervals [of f "interior x"]
-        apply (auto simp: insert)
-        by (metis IntI empty_iff insert.hyps(2) insert.prems(3) insert_iff)
-    qed
-  } note UN_cases = this
-  let ?A = "{cbox c d | c d::'a. \<forall>i\<in>Basis. (c\<bullet>i = a\<bullet>i) \<and> (d\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<or>
-    (c\<bullet>i = (a\<bullet>i + b\<bullet>i) / 2) \<and> (d\<bullet>i = b\<bullet>i)}"
-  let ?PP = "\<lambda>c d. \<forall>i\<in>Basis. a\<bullet>i \<le> c\<bullet>i \<and> c\<bullet>i \<le> d\<bullet>i \<and> d\<bullet>i \<le> b\<bullet>i \<and> 2 * (d\<bullet>i - c\<bullet>i) \<le> b\<bullet>i - a\<bullet>i"
-  {
-    presume "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d) \<Longrightarrow> False"
-    then show thesis
-      unfolding atomize_not not_all
-      by (blast intro: that)
-  }
-  assume as: "\<forall>c d. ?PP c d \<longrightarrow> P (cbox c d)"
-  have "P (\<Union>?A)"
-  proof (rule UN_cases)
-    let ?B = "(\<lambda>s. cbox (\<Sum>i\<in>Basis. (if i \<in> s then a\<bullet>i else (a\<bullet>i + b\<bullet>i) / 2) *\<^sub>R i::'a)
-      (\<Sum>i\<in>Basis. (if i \<in> s then (a\<bullet>i + b\<bullet>i) / 2 else b\<bullet>i) *\<^sub>R i)) ` {s. s \<subseteq> Basis}"
-    have "?A \<subseteq> ?B"
-    proof
-      fix x
-      assume "x \<in> ?A"
-      then obtain c d
-        where x:  "x = cbox c d"
-                  "\<And>i. i \<in> Basis \<Longrightarrow>
-                        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-                        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i" by blast
-      show "x \<in> ?B"
-        unfolding image_iff x
-        apply (rule_tac x="{i. i\<in>Basis \<and> c\<bullet>i = a\<bullet>i}" in bexI)
-        apply (rule arg_cong2 [where f = cbox])
-        using x(2) ab
-        apply (auto simp add: euclidean_eq_iff[where 'a='a])
-        by fastforce
-    qed
-    then show "finite ?A"
-      by (rule finite_subset) auto
-  next
-    fix s
-    assume "s \<in> ?A"
-    then obtain c d
-      where s: "s = cbox c d"
-               "\<And>i. i \<in> Basis \<Longrightarrow>
-                     c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-                     c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
-      by blast
-    show "P s"
-      unfolding s
-      apply (rule as[rule_format])
-      using ab s(2) by force
-    show "\<exists>a b. s = cbox a b"
-      unfolding s by auto
-    fix t
-    assume "t \<in> ?A"
-    then obtain e f where t:
-      "t = cbox e f"
-      "\<And>i. i \<in> Basis \<Longrightarrow>
-        e \<bullet> i = a \<bullet> i \<and> f \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-        e \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> f \<bullet> i = b \<bullet> i"
-      by blast
-    assume "s \<noteq> t"
-    then have "\<not> (c = e \<and> d = f)"
-      unfolding s t by auto
-    then obtain i where "c\<bullet>i \<noteq> e\<bullet>i \<or> d\<bullet>i \<noteq> f\<bullet>i" and i': "i \<in> Basis"
-      unfolding euclidean_eq_iff[where 'a='a] by auto
-    then have i: "c\<bullet>i \<noteq> e\<bullet>i" "d\<bullet>i \<noteq> f\<bullet>i"
-      using s(2) t(2) apply fastforce
-      using t(2)[OF i'] \<open>c \<bullet> i \<noteq> e \<bullet> i \<or> d \<bullet> i \<noteq> f \<bullet> i\<close> i' s(2) t(2) by fastforce
-    have *: "\<And>s t. (\<And>a. a \<in> s \<Longrightarrow> a \<in> t \<Longrightarrow> False) \<Longrightarrow> s \<inter> t = {}"
-      by auto
-    show "interior s \<inter> interior t = {}"
-      unfolding s t interior_cbox
-    proof (rule *)
-      fix x
-      assume "x \<in> box c d" "x \<in> box e f"
-      then have x: "c\<bullet>i < d\<bullet>i" "e\<bullet>i < f\<bullet>i" "c\<bullet>i < f\<bullet>i" "e\<bullet>i < d\<bullet>i"
-        unfolding mem_box using i'
-        by force+
-      show False  using s(2)[OF i']
-      proof safe
-        assume as: "c \<bullet> i = a \<bullet> i" "d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2"
-        show False
-          using t(2)[OF i'] and i x unfolding as by (fastforce simp add:field_simps)
-      next
-        assume as: "c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2" "d \<bullet> i = b \<bullet> i"
-        show False
-          using t(2)[OF i'] and i x unfolding as by(fastforce simp add:field_simps)
-      qed
-    qed
-  qed
-  also have "\<Union>?A = cbox a b"
-  proof (rule set_eqI,rule)
-    fix x
-    assume "x \<in> \<Union>?A"
-    then obtain c d where x:
-      "x \<in> cbox c d"
-      "\<And>i. i \<in> Basis \<Longrightarrow>
-        c \<bullet> i = a \<bullet> i \<and> d \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<or>
-        c \<bullet> i = (a \<bullet> i + b \<bullet> i) / 2 \<and> d \<bullet> i = b \<bullet> i"
-      by blast
-    show "x\<in>cbox a b"
-      unfolding mem_box
-    proof safe
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      then show "a \<bullet> i \<le> x \<bullet> i" "x \<bullet> i \<le> b \<bullet> i"
-        using x(2)[OF i] x(1)[unfolded mem_box,THEN bspec, OF i] by auto
-    qed
-  next
-    fix x
-    assume x: "x \<in> cbox a b"
-    have "\<forall>i\<in>Basis.
-      \<exists>c d. (c = a\<bullet>i \<and> d = (a\<bullet>i + b\<bullet>i) / 2 \<or> c = (a\<bullet>i + b\<bullet>i) / 2 \<and> d = b\<bullet>i) \<and> c\<le>x\<bullet>i \<and> x\<bullet>i \<le> d"
-      (is "\<forall>i\<in>Basis. \<exists>c d. ?P i c d")
-      unfolding mem_box
-    proof
-      fix i :: 'a
-      assume i: "i \<in> Basis"
-      have "?P i (a\<bullet>i) ((a \<bullet> i + b \<bullet> i) / 2) \<or> ?P i ((a \<bullet> i + b \<bullet> i) / 2) (b\<bullet>i)"
-        using x[unfolded mem_box,THEN bspec, OF i] by auto
-      then show "\<exists>c d. ?P i c d"
-        by blast
-    qed
-    then show "x\<in>\<Union>?A"
-      unfolding Union_iff Bex_def mem_Collect_eq choice_Basis_iff
-      apply auto
-      apply (rule_tac x="cbox xa xaa" in exI)
-      unfolding mem_box
-      apply auto
-      done
-  qed
-  finally show False
-    using assms by auto
-qed
-
-lemma interval_bisection:
-  fixes type :: "'a::euclidean_space"
-  assumes "P {}"
-    and "(\<forall>s t. P s \<and> P t \<and> interior(s) \<inter> interior(t) = {} \<longrightarrow> P(s \<union> t))"
-    and "\<not> P (cbox a (b::'a))"
-  obtains x where "x \<in> cbox a b"
-    and "\<forall>e>0. \<exists>c d. x \<in> cbox c d \<and> cbox c d \<subseteq> ball x e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-proof -
-  have "\<forall>x. \<exists>y. \<not> P (cbox (fst x) (snd x)) \<longrightarrow> (\<not> P (cbox (fst y) (snd y)) \<and>
-    (\<forall>i\<in>Basis. fst x\<bullet>i \<le> fst y\<bullet>i \<and> fst y\<bullet>i \<le> snd y\<bullet>i \<and> snd y\<bullet>i \<le> snd x\<bullet>i \<and>
-       2 * (snd y\<bullet>i - fst y\<bullet>i) \<le> snd x\<bullet>i - fst x\<bullet>i))" (is "\<forall>x. ?P x")
-  proof
-    show "?P x" for x
-    proof (cases "P (cbox (fst x) (snd x))")
-      case True
-      then show ?thesis by auto
-    next
-      case as: False
-      obtain c d where "\<not> P (cbox c d)"
-        "\<forall>i\<in>Basis.
-           fst x \<bullet> i \<le> c \<bullet> i \<and>
-           c \<bullet> i \<le> d \<bullet> i \<and>
-           d \<bullet> i \<le> snd x \<bullet> i \<and>
-           2 * (d \<bullet> i - c \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i"
-        by (rule interval_bisection_step[of P, OF assms(1-2) as])
-      then show ?thesis
-        apply -
-        apply (rule_tac x="(c,d)" in exI)
-        apply auto
-        done
-    qed
-  qed
-  then obtain f where f:
-    "\<forall>x.
-      \<not> P (cbox (fst x) (snd x)) \<longrightarrow>
-      \<not> P (cbox (fst (f x)) (snd (f x))) \<and>
-        (\<forall>i\<in>Basis.
-            fst x \<bullet> i \<le> fst (f x) \<bullet> i \<and>
-            fst (f x) \<bullet> i \<le> snd (f x) \<bullet> i \<and>
-            snd (f x) \<bullet> i \<le> snd x \<bullet> i \<and>
-            2 * (snd (f x) \<bullet> i - fst (f x) \<bullet> i) \<le> snd x \<bullet> i - fst x \<bullet> i)"
-    apply -
-    apply (drule choice)
-    apply blast
-    done
-  define AB A B where ab_def: "AB n = (f ^^ n) (a,b)" "A n = fst(AB n)" "B n = snd(AB n)" for n
-  have "A 0 = a" "B 0 = b" "\<And>n. \<not> P (cbox (A(Suc n)) (B(Suc n))) \<and>
-    (\<forall>i\<in>Basis. A(n)\<bullet>i \<le> A(Suc n)\<bullet>i \<and> A(Suc n)\<bullet>i \<le> B(Suc n)\<bullet>i \<and> B(Suc n)\<bullet>i \<le> B(n)\<bullet>i \<and>
-    2 * (B(Suc n)\<bullet>i - A(Suc n)\<bullet>i) \<le> B(n)\<bullet>i - A(n)\<bullet>i)" (is "\<And>n. ?P n")
-  proof -
-    show "A 0 = a" "B 0 = b"
-      unfolding ab_def by auto
-    note S = ab_def funpow.simps o_def id_apply
-    show "?P n" for n
-    proof (induct n)
-      case 0
-      then show ?case
-        unfolding S
-        apply (rule f[rule_format]) using assms(3)
-        apply auto
-        done
-    next
-      case (Suc n)
-      show ?case
-        unfolding S
-        apply (rule f[rule_format])
-        using Suc
-        unfolding S
-        apply auto
-        done
-    qed
-  qed
-  note AB = this(1-2) conjunctD2[OF this(3),rule_format]
-
-  have interv: "\<exists>n. \<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e"
-    if e: "0 < e" for e
-  proof -
-    obtain n where n: "(\<Sum>i\<in>Basis. b \<bullet> i - a \<bullet> i) / e < 2 ^ n"
-      using real_arch_pow[of 2 "(setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis) / e"] by auto
-    show ?thesis
-    proof (rule exI [where x=n], clarify)
-      fix x y
-      assume xy: "x\<in>cbox (A n) (B n)" "y\<in>cbox (A n) (B n)"
-      have "dist x y \<le> setsum (\<lambda>i. \<bar>(x - y)\<bullet>i\<bar>) Basis"
-        unfolding dist_norm by(rule norm_le_l1)
-      also have "\<dots> \<le> setsum (\<lambda>i. B n\<bullet>i - A n\<bullet>i) Basis"
-      proof (rule setsum_mono)
-        fix i :: 'a
-        assume i: "i \<in> Basis"
-        show "\<bar>(x - y) \<bullet> i\<bar> \<le> B n \<bullet> i - A n \<bullet> i"
-          using xy[unfolded mem_box,THEN bspec, OF i]
-          by (auto simp: inner_diff_left)
-      qed
-      also have "\<dots> \<le> setsum (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis / 2^n"
-        unfolding setsum_divide_distrib
-      proof (rule setsum_mono)
-        show "B n \<bullet> i - A n \<bullet> i \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ n" if i: "i \<in> Basis" for i
-        proof (induct n)
-          case 0
-          then show ?case
-            unfolding AB by auto
-        next
-          case (Suc n)
-          have "B (Suc n) \<bullet> i - A (Suc n) \<bullet> i \<le> (B n \<bullet> i - A n \<bullet> i) / 2"
-            using AB(4)[of i n] using i by auto
-          also have "\<dots> \<le> (b \<bullet> i - a \<bullet> i) / 2 ^ Suc n"
-            using Suc by (auto simp add: field_simps)
-          finally show ?case .
-        qed
-      qed
-      also have "\<dots> < e"
-        using n using e by (auto simp add: field_simps)
-      finally show "dist x y < e" .
-    qed
-  qed
-  {
-    fix n m :: nat
-    assume "m \<le> n" then have "cbox (A n) (B n) \<subseteq> cbox (A m) (B m)"
-    proof (induction rule: inc_induct)
-      case (step i)
-      show ?case
-        using AB(4) by (intro order_trans[OF step.IH] subset_box_imp) auto
-    qed simp
-  } note ABsubset = this
-  have "\<exists>a. \<forall>n. a\<in> cbox (A n) (B n)"
-    by (rule decreasing_closed_nest[rule_format,OF closed_cbox _ ABsubset interv])
-      (metis nat.exhaust AB(1-3) assms(1,3))
-  then obtain x0 where x0: "\<And>n. x0 \<in> cbox (A n) (B n)"
-    by blast
-  show thesis
-  proof (rule that[rule_format, of x0])
-    show "x0\<in>cbox a b"
-      using x0[of 0] unfolding AB .
-    fix e :: real
-    assume "e > 0"
-    from interv[OF this] obtain n
-      where n: "\<forall>x\<in>cbox (A n) (B n). \<forall>y\<in>cbox (A n) (B n). dist x y < e" ..
-    have "\<not> P (cbox (A n) (B n))"
-      apply (cases "0 < n")
-      using AB(3)[of "n - 1"] assms(3) AB(1-2)
-      apply auto
-      done
-    moreover have "cbox (A n) (B n) \<subseteq> ball x0 e"
-      using n using x0[of n] by auto
-    moreover have "cbox (A n) (B n) \<subseteq> cbox a b"
-      unfolding AB(1-2)[symmetric] by (rule ABsubset) auto
-    ultimately show "\<exists>c d. x0 \<in> cbox c d \<and> cbox c d \<subseteq> ball x0 e \<and> cbox c d \<subseteq> cbox a b \<and> \<not> P (cbox c d)"
-      apply (rule_tac x="A n" in exI)
-      apply (rule_tac x="B n" in exI)
-      apply (auto simp: x0)
-      done
-  qed
-qed
-
-
-subsection \<open>Cousin's lemma.\<close>
-
-lemma fine_division_exists:
-  fixes a b :: "'a::euclidean_space"
-  assumes "gauge g"
-  obtains p where "p tagged_division_of (cbox a b)" "g fine p"
-proof -
-  presume "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p) \<Longrightarrow> False"
-  then obtain p where "p tagged_division_of (cbox a b)" "g fine p"
-    by blast
-  then show thesis ..
-next
-  assume as: "\<not> (\<exists>p. p tagged_division_of (cbox a b) \<and> g fine p)"
-  obtain x where x:
-      "x \<in> (cbox a b)"
-      "\<And>e. 0 < e \<Longrightarrow>
-        \<exists>c d.
-          x \<in> cbox c d \<and>
-          cbox c d \<subseteq> ball x e \<and>
-          cbox c d \<subseteq> (cbox a b) \<and>
-          \<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    apply (rule interval_bisection[of "\<lambda>s. \<exists>p. p tagged_division_of s \<and> g fine p", OF _ _ as])
-    apply (simp add: fine_def)
-    apply (metis tagged_division_union fine_union)
-    apply (auto simp: )
-    done
-  obtain e where e: "e > 0" "ball x e \<subseteq> g x"
-    using gaugeD[OF assms, of x] unfolding open_contains_ball by auto
-  from x(2)[OF e(1)]
-  obtain c d where c_d: "x \<in> cbox c d"
-                        "cbox c d \<subseteq> ball x e"
-                        "cbox c d \<subseteq> cbox a b"
-                        "\<not> (\<exists>p. p tagged_division_of cbox c d \<and> g fine p)"
-    by blast
-  have "g fine {(x, cbox c d)}"
-    unfolding fine_def using e using c_d(2) by auto
-  then show False
-    using tagged_division_of_self[OF c_d(1)] using c_d by auto
-qed
-
-lemma fine_division_exists_real:
-  fixes a b :: real
-  assumes "gauge g"
-  obtains p where "p tagged_division_of {a .. b}" "g fine p"
-  by (metis assms box_real(2) fine_division_exists)
-
-subsection \<open>Basic theorems about integrals.\<close>
-
-lemma has_integral_unique:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral k1) i"
-    and "(f has_integral k2) i"
-  shows "k1 = k2"
-proof (rule ccontr)
-  let ?e = "norm (k1 - k2) / 2"
-  assume as: "k1 \<noteq> k2"
-  then have e: "?e > 0"
-    by auto
-  have lem: False
-    if f_k1: "(f has_integral k1) (cbox a b)"
-    and f_k2: "(f has_integral k2) (cbox a b)"
-    and "k1 \<noteq> k2"
-    for f :: "'n \<Rightarrow> 'a" and a b k1 k2
-  proof -
-    let ?e = "norm (k1 - k2) / 2"
-    from \<open>k1 \<noteq> k2\<close> have e: "?e > 0" by auto
-    obtain d1 where d1:
-        "gauge d1"
-        "\<And>p. p tagged_division_of cbox a b \<Longrightarrow>
-          d1 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k1) < norm (k1 - k2) / 2"
-      by (rule has_integralD[OF f_k1 e]) blast
-    obtain d2 where d2:
-        "gauge d2"
-        "\<And>p. p tagged_division_of cbox a b \<Longrightarrow>
-          d2 fine p \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k2) < norm (k1 - k2) / 2"
-      by (rule has_integralD[OF f_k2 e]) blast
-    obtain p where p:
-        "p tagged_division_of cbox a b"
-        "(\<lambda>x. d1 x \<inter> d2 x) fine p"
-      by (rule fine_division_exists[OF gauge_inter[OF d1(1) d2(1)]])
-    let ?c = "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
-    have "norm (k1 - k2) \<le> norm (?c - k2) + norm (?c - k1)"
-      using norm_triangle_ineq4[of "k1 - ?c" "k2 - ?c"]
-      by (auto simp add:algebra_simps norm_minus_commute)
-    also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
-      apply (rule add_strict_mono)
-      apply (rule_tac[!] d2(2) d1(2))
-      using p unfolding fine_def
-      apply auto
-      done
-    finally show False by auto
-  qed
-  {
-    presume "\<not> (\<exists>a b. i = cbox a b) \<Longrightarrow> False"
-    then show False
-      using as assms lem by blast
-  }
-  assume as: "\<not> (\<exists>a b. i = cbox a b)"
-  obtain B1 where B1:
-      "0 < B1"
-      "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow>
-        \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and>
-          norm (z - k1) < norm (k1 - k2) / 2"
-    by (rule has_integral_altD[OF assms(1) as,OF e]) blast
-  obtain B2 where B2:
-      "0 < B2"
-      "\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow>
-        \<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and>
-          norm (z - k2) < norm (k1 - k2) / 2"
-    by (rule has_integral_altD[OF assms(2) as,OF e]) blast
-  have "\<exists>a b::'n. ball 0 B1 \<union> ball 0 B2 \<subseteq> cbox a b"
-    apply (rule bounded_subset_cbox)
-    using bounded_Un bounded_ball
-    apply auto
-    done
-  then obtain a b :: 'n where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b"
-    by blast
-  obtain w where w:
-    "((\<lambda>x. if x \<in> i then f x else 0) has_integral w) (cbox a b)"
-    "norm (w - k1) < norm (k1 - k2) / 2"
-    using B1(2)[OF ab(1)] by blast
-  obtain z where z:
-    "((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b)"
-    "norm (z - k2) < norm (k1 - k2) / 2"
-    using B2(2)[OF ab(2)] by blast
-  have "z = w"
-    using lem[OF w(1) z(1)] by auto
-  then have "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)"
-    using norm_triangle_ineq4 [of "k1 - w" "k2 - z"]
-    by (auto simp add: norm_minus_commute)
-  also have "\<dots> < norm (k1 - k2) / 2 + norm (k1 - k2) / 2"
-    apply (rule add_strict_mono)
-    apply (rule_tac[!] z(2) w(2))
-    done
-  finally show False by auto
-qed
-
-lemma integral_unique [intro]: "(f has_integral y) k \<Longrightarrow> integral k f = y"
-  unfolding integral_def
-  by (rule some_equality) (auto intro: has_integral_unique)
-
-lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> ~ f integrable_on k \<and> y=0"
-  unfolding integral_def integrable_on_def
-  apply (erule subst)
-  apply (rule someI_ex)
-  by blast
-
-lemma has_integral_is_0:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "\<forall>x\<in>s. f x = 0"
-  shows "(f has_integral 0) s"
-proof -
-  have lem: "\<And>a b. \<And>f::'n \<Rightarrow> 'a.
-    (\<forall>x\<in>cbox a b. f(x) = 0) \<Longrightarrow> (f has_integral 0) (cbox a b)"
-    unfolding has_integral
-  proof clarify
-    fix a b e
-    fix f :: "'n \<Rightarrow> 'a"
-    assume as: "\<forall>x\<in>cbox a b. f x = 0" "0 < (e::real)"
-    have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
-      if p: "p tagged_division_of cbox a b" for p
-    proof -
-      have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) = 0"
-      proof (rule setsum.neutral, rule)
-        fix x
-        assume x: "x \<in> p"
-        have "f (fst x) = 0"
-          using tagged_division_ofD(2-3)[OF p, of "fst x" "snd x"] using as x by auto
-        then show "(\<lambda>(x, k). content k *\<^sub>R f x) x = 0"
-          apply (subst surjective_pairing[of x])
-          unfolding split_conv
-          apply auto
-          done
-      qed
-      then show ?thesis
-        using as by auto
-    qed
-    then show "\<exists>d. gauge d \<and>
-        (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e)"
-      by auto
-  qed
-  {
-    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
-    with assms lem show ?thesis
-      by blast
-  }
-  have *: "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. 0)"
-    apply (rule ext)
-    using assms
-    apply auto
-    done
-  assume "\<not> (\<exists>a b. s = cbox a b)"
-  then show ?thesis
-    using lem
-    by (subst has_integral_alt) (force simp add: *)
-qed
-
-lemma has_integral_0[simp]: "((\<lambda>x::'n::euclidean_space. 0) has_integral 0) s"
-  by (rule has_integral_is_0) auto
-
-lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) s \<longleftrightarrow> i = 0"
-  using has_integral_unique[OF has_integral_0] by auto
-
-lemma has_integral_linear:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral y) s"
-    and "bounded_linear h"
-  shows "((h \<circ> f) has_integral ((h y))) s"
-proof -
-  interpret bounded_linear h
-    using assms(2) .
-  from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
-    by blast
-  have lem: "\<And>(f :: 'n \<Rightarrow> 'a) y a b.
-    (f has_integral y) (cbox a b) \<Longrightarrow> ((h \<circ> f) has_integral h y) (cbox a b)"
-    unfolding has_integral
-  proof (clarify, goal_cases)
-    case prems: (1 f y a b e)
-    from pos_bounded
-    obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B"
-      by blast
-    have "e / B > 0" using prems(2) B by simp
-    then obtain g
-      where g: "gauge g"
-               "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> g fine p \<Longrightarrow>
-                    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e / B"
-        using prems(1) by auto
-    {
-      fix p
-      assume as: "p tagged_division_of (cbox a b)" "g fine p"
-      have hc: "\<And>x k. h ((\<lambda>(x, k). content k *\<^sub>R f x) x) = (\<lambda>(x, k). h (content k *\<^sub>R f x)) x"
-        by auto
-      then have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = setsum (h \<circ> (\<lambda>(x, k). content k *\<^sub>R f x)) p"
-        unfolding o_def unfolding scaleR[symmetric] hc by simp
-      also have "\<dots> = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
-        using setsum[of "\<lambda>(x,k). content k *\<^sub>R f x" p] using as by auto
-      finally have "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) = h (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)" .
-      then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (h \<circ> f) x) - h y) < e"
-        apply (simp add: diff[symmetric])
-        apply (rule le_less_trans[OF B(2)])
-        using g(2)[OF as] B(1)
-        apply (auto simp add: field_simps)
-        done
-    }
-    with g show ?case
-      by (rule_tac x=g in exI) auto
-  qed
-  {
-    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
-    then show ?thesis
-      using assms(1) lem by blast
-  }
-  assume as: "\<not> (\<exists>a b. s = cbox a b)"
-  then show ?thesis
-  proof (subst has_integral_alt, clarsimp)
-    fix e :: real
-    assume e: "e > 0"
-    have *: "0 < e/B" using e B(1) by simp
-    obtain M where M:
-      "M > 0"
-      "\<And>a b. ball 0 M \<subseteq> cbox a b \<Longrightarrow>
-        \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e / B"
-      using has_integral_altD[OF assms(1) as *] by blast
-    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) has_integral z) (cbox a b) \<and> norm (z - h y) < e)"
-    proof (rule_tac x=M in exI, clarsimp simp add: M, goal_cases)
-      case prems: (1 a b)
-      obtain z where z:
-        "((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b)"
-        "norm (z - y) < e / B"
-        using M(2)[OF prems(1)] by blast
-      have *: "(\<lambda>x. if x \<in> s then (h \<circ> f) x else 0) = h \<circ> (\<lambda>x. if x \<in> s then f x else 0)"
-        using zero by auto
-      show ?case
-        apply (rule_tac x="h z" in exI)
-        apply (simp add: * lem z(1))
-        apply (metis B diff le_less_trans pos_less_divide_eq z(2))
-        done
-    qed
-  qed
-qed
-
-lemma has_integral_scaleR_left:
-  "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s"
-  using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def)
-
-lemma has_integral_mult_left:
-  fixes c :: "_ :: real_normed_algebra"
-  shows "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) s"
-  using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def)
-
-text\<open>The case analysis eliminates the condition @{term "f integrable_on s"} at the cost
-     of the type class constraint \<open>division_ring\<close>\<close>
-corollary integral_mult_left [simp]:
-  fixes c:: "'a::{real_normed_algebra,division_ring}"
-  shows "integral s (\<lambda>x. f x * c) = integral s f * c"
-proof (cases "f integrable_on s \<or> c = 0")
-  case True then show ?thesis
-    by (force intro: has_integral_mult_left)
-next
-  case False then have "~ (\<lambda>x. f x * c) integrable_on s"
-    using has_integral_mult_left [of "(\<lambda>x. f x * c)" _ s "inverse c"]
-    by (force simp add: mult.assoc)
-  with False show ?thesis by (simp add: not_integrable_integral)
-qed
-
-corollary integral_mult_right [simp]:
-  fixes c:: "'a::{real_normed_field}"
-  shows "integral s (\<lambda>x. c * f x) = c * integral s f"
-by (simp add: mult.commute [of c])
-
-corollary integral_divide [simp]:
-  fixes z :: "'a::real_normed_field"
-  shows "integral S (\<lambda>x. f x / z) = integral S (\<lambda>x. f x) / z"
-using integral_mult_left [of S f "inverse z"]
-  by (simp add: divide_inverse_commute)
-
-lemma has_integral_mult_right:
-  fixes c :: "'a :: real_normed_algebra"
-  shows "(f has_integral y) i \<Longrightarrow> ((\<lambda>x. c * f x) has_integral (c * y)) i"
-  using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def)
-
-lemma has_integral_cmul: "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) s"
-  unfolding o_def[symmetric]
-  by (metis has_integral_linear bounded_linear_scaleR_right)
-
-lemma has_integral_cmult_real:
-  fixes c :: real
-  assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A"
-  shows "((\<lambda>x. c * f x) has_integral c * x) A"
-proof (cases "c = 0")
-  case True
-  then show ?thesis by simp
-next
-  case False
-  from has_integral_cmul[OF assms[OF this], of c] show ?thesis
-    unfolding real_scaleR_def .
-qed
-
-lemma has_integral_neg: "(f has_integral k) s \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral -k) s"
-  by (drule_tac c="-1" in has_integral_cmul) auto
-
-lemma has_integral_add:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "(f has_integral k) s"
-    and "(g has_integral l) s"
-  shows "((\<lambda>x. f x + g x) has_integral (k + l)) s"
-proof -
-  have lem: "((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)"
-    if f_k: "(f has_integral k) (cbox a b)"
-    and g_l: "(g has_integral l) (cbox a b)"
-    for f :: "'n \<Rightarrow> 'a" and g a b k l
-    unfolding has_integral
-  proof clarify
-    fix e :: real
-    assume e: "e > 0"
-    then have *: "e / 2 > 0"
-      by auto
-    obtain d1 where d1:
-      "gauge d1"
-      "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d1 fine p \<Longrightarrow>
-        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) < e / 2"
-      using has_integralD[OF f_k *] by blast
-    obtain d2 where d2:
-      "gauge d2"
-      "\<And>p. p tagged_division_of (cbox a b) \<Longrightarrow> d2 fine p \<Longrightarrow>
-        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l) < e / 2"
-      using has_integralD[OF g_l *] by blast
-    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-              norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e)"
-    proof (rule exI [where x="\<lambda>x. (d1 x) \<inter> (d2 x)"], clarsimp simp add: gauge_inter[OF d1(1) d2(1)])
-      fix p
-      assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d1 x \<inter> d2 x) fine p"
-      have *: "(\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) =
-        (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p. content k *\<^sub>R g x)"
-        unfolding scaleR_right_distrib setsum.distrib[of "\<lambda>(x,k). content k *\<^sub>R f x" "\<lambda>(x,k). content k *\<^sub>R g x" p,symmetric]
-        by (rule setsum.cong) auto
-      from as have fine: "d1 fine p" "d2 fine p"
-        unfolding fine_inter by auto
-      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) =
-            norm (((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - k) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - l))"
-        unfolding * by (auto simp add: algebra_simps)
-      also have "\<dots> < e/2 + e/2"
-        apply (rule le_less_trans[OF norm_triangle_ineq])
-        using as d1 d2 fine
-        apply (blast intro: add_strict_mono)
-        done
-      finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R (f x + g x)) - (k + l)) < e"
-        by auto
-    qed
-  qed
-  {
-    presume "\<not> (\<exists>a b. s = cbox a b) \<Longrightarrow> ?thesis"
-    then show ?thesis
-      using assms lem by force
-  }
-  assume as: "\<not> (\<exists>a b. s = cbox a b)"
-  then show ?thesis
-  proof (subst has_integral_alt, clarsimp, goal_cases)
-    case (1 e)
-    then have *: "e / 2 > 0"
-      by auto
-    from has_integral_altD[OF assms(1) as *]
-    obtain B1 where B1:
-        "0 < B1"
-        "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow>
-          \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - k) < e / 2"
-      by blast
-    from has_integral_altD[OF assms(2) as *]
-    obtain B2 where B2:
-        "0 < B2"
-        "\<And>a b. ball 0 B2 \<subseteq> (cbox a b) \<Longrightarrow>
-          \<exists>z. ((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b) \<and> norm (z - l) < e / 2"
-      by blast
-    show ?case
-    proof (rule_tac x="max B1 B2" in exI, clarsimp simp add: max.strict_coboundedI1 B1)
-      fix a b
-      assume "ball 0 (max B1 B2) \<subseteq> cbox a (b::'n)"
-      then have *: "ball 0 B1 \<subseteq> cbox a (b::'n)" "ball 0 B2 \<subseteq> cbox a (b::'n)"
-        by auto
-      obtain w where w:
-        "((\<lambda>x. if x \<in> s then f x else 0) has_integral w) (cbox a b)"
-        "norm (w - k) < e / 2"
-        using B1(2)[OF *(1)] by blast
-      obtain z where z:
-        "((\<lambda>x. if x \<in> s then g x else 0) has_integral z) (cbox a b)"
-        "norm (z - l) < e / 2"
-        using B2(2)[OF *(2)] by blast
-      have *: "\<And>x. (if x \<in> s then f x + g x else 0) =
-        (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0)"
-        by auto
-      show "\<exists>z. ((\<lambda>x. if x \<in> s then f x + g x else 0) has_integral z) (cbox a b) \<and> norm (z - (k + l)) < e"
-        apply (rule_tac x="w + z" in exI)
-        apply (simp add: lem[OF w(1) z(1), unfolded *[symmetric]])
-        using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2)
-        apply (auto simp add: field_simps)
-        done
-    qed
-  qed
-qed
-
-lemma has_integral_sub:
-  "(f has_integral k) s \<Longrightarrow> (g has_integral l) s \<Longrightarrow>
-    ((\<lambda>x. f x - g x) has_integral (k - l)) s"
-  using has_integral_add[OF _ has_integral_neg, of f k s g l]
-  by (auto simp: algebra_simps)
-
-lemma integral_0 [simp]:
-  "integral s (\<lambda>x::'n::euclidean_space. 0::'m::real_normed_vector) = 0"
-  by (rule integral_unique has_integral_0)+
-
-lemma integral_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
-    integral s (\<lambda>x. f x + g x) = integral s f + integral s g"
-  by (rule integral_unique) (metis integrable_integral has_integral_add)
-
-lemma integral_cmul [simp]: "integral s (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral s f"
-proof (cases "f integrable_on s \<or> c = 0")
-  case True with has_integral_cmul show ?thesis by force
-next
-  case False then have "~ (\<lambda>x. c *\<^sub>R f x) integrable_on s"
-    using has_integral_cmul [of "(\<lambda>x. c *\<^sub>R f x)" _ s "inverse c"]
-    by force
-  with False show ?thesis by (simp add: not_integrable_integral)
-qed
-
-lemma integral_neg [simp]: "integral s (\<lambda>x. - f x) = - integral s f"
-proof (cases "f integrable_on s")
-  case True then show ?thesis
-    by (simp add: has_integral_neg integrable_integral integral_unique)
-next
-  case False then have "~ (\<lambda>x. - f x) integrable_on s"
-    using has_integral_neg [of "(\<lambda>x. - f x)" _ s ]
-    by force
-  with False show ?thesis by (simp add: not_integrable_integral)
-qed
-
-lemma integral_diff: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow>
-    integral s (\<lambda>x. f x - g x) = integral s f - integral s g"
-  by (rule integral_unique) (metis integrable_integral has_integral_sub)
-
-lemma integrable_0: "(\<lambda>x. 0) integrable_on s"
-  unfolding integrable_on_def using has_integral_0 by auto
-
-lemma integrable_add: "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x + g x) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_add)
-
-lemma integrable_cmul: "f integrable_on s \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_cmul)
-
-lemma integrable_on_cmult_iff:
-  fixes c :: real
-  assumes "c \<noteq> 0"
-  shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
-  using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] \<open>c \<noteq> 0\<close>
-  by auto
-
-lemma integrable_on_cmult_left:
-  assumes "f integrable_on s"
-  shows "(\<lambda>x. of_real c * f x) integrable_on s"
-    using integrable_cmul[of f s "of_real c"] assms
-    by (simp add: scaleR_conv_of_real)
-
-lemma integrable_neg: "f integrable_on s \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_neg)
-
-lemma integrable_diff:
-  "f integrable_on s \<Longrightarrow> g integrable_on s \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_sub)
-
-lemma integrable_linear:
-  "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) integrable_on s"
-  unfolding integrable_on_def by(auto intro: has_integral_linear)
-
-lemma integral_linear:
-  "f integrable_on s \<Longrightarrow> bounded_linear h \<Longrightarrow> integral s (h \<circ> f) = h (integral s f)"
-  apply (rule has_integral_unique [where i=s and f = "h \<circ> f"])
-  apply (simp_all add: integrable_integral integrable_linear has_integral_linear )
-  done
-
-lemma integral_component_eq[simp]:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "f integrable_on s"
-  shows "integral s (\<lambda>x. f x \<bullet> k) = integral s f \<bullet> k"
-  unfolding integral_linear[OF assms(1) bounded_linear_component,unfolded o_def] ..
-
-lemma has_integral_setsum:
-  assumes "finite t"
-    and "\<forall>a\<in>t. ((f a) has_integral (i a)) s"
-  shows "((\<lambda>x. setsum (\<lambda>a. f a x) t) has_integral (setsum i t)) s"
-  using assms(1) subset_refl[of t]
-proof (induct rule: finite_subset_induct)
-  case empty
-  then show ?case by auto
-next
-  case (insert x F)
-  with assms show ?case
-    by (simp add: has_integral_add)
-qed
-
-lemma integral_setsum:
-  "\<lbrakk>finite t;  \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow>
-   integral s (\<lambda>x. setsum (\<lambda>a. f a x) t) = setsum (\<lambda>a. integral s (f a)) t"
-  by (auto intro: has_integral_setsum integrable_integral)
-
-lemma integrable_setsum:
-  "\<lbrakk>finite t;  \<forall>a\<in>t. (f a) integrable_on s\<rbrakk> \<Longrightarrow> (\<lambda>x. setsum (\<lambda>a. f a x) t) integrable_on s"
-  unfolding integrable_on_def
-  apply (drule bchoice)
-  using has_integral_setsum[of t]
-  apply auto
-  done
-
-lemma has_integral_eq:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-    and "(f has_integral k) s"
-  shows "(g has_integral k) s"
-  using has_integral_sub[OF assms(2), of "\<lambda>x. f x - g x" 0]
-  using has_integral_is_0[of s "\<lambda>x. f x - g x"]
-  using assms(1)
-  by auto
-
-lemma integrable_eq: "(\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f integrable_on s \<Longrightarrow> g integrable_on s"
-  unfolding integrable_on_def
-  using has_integral_eq[of s f g] has_integral_eq by blast
-
-lemma has_integral_cong:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-  shows "(f has_integral i) s = (g has_integral i) s"
-  using has_integral_eq[of s f g] has_integral_eq[of s g f] assms
-  by auto
-
-lemma integral_cong:
-  assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-  shows "integral s f = integral s g"
-  unfolding integral_def
-by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq)
-
-lemma integrable_on_cmult_left_iff [simp]:
-  assumes "c \<noteq> 0"
-  shows "(\<lambda>x. of_real c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
-        (is "?lhs = ?rhs")
-proof
-  assume ?lhs
-  then have "(\<lambda>x. of_real (1 / c) * (of_real c * f x)) integrable_on s"
-    using integrable_cmul[of "\<lambda>x. of_real c * f x" s "1 / of_real c"]
-    by (simp add: scaleR_conv_of_real)
-  then have "(\<lambda>x. (of_real (1 / c) * of_real c * f x)) integrable_on s"
-    by (simp add: algebra_simps)
-  with \<open>c \<noteq> 0\<close> show ?rhs
-    by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult)
-qed (blast intro: integrable_on_cmult_left)
-
-lemma integrable_on_cmult_right:
-  fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
-  assumes "f integrable_on s"
-  shows "(\<lambda>x. f x * of_real c) integrable_on s"
-using integrable_on_cmult_left [OF assms] by (simp add: mult.commute)
-
-lemma integrable_on_cmult_right_iff [simp]:
-  fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
-  assumes "c \<noteq> 0"
-  shows "(\<lambda>x. f x * of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
-using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute)
-
-lemma integrable_on_cdivide:
-  fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
-  assumes "f integrable_on s"
-  shows "(\<lambda>x. f x / of_real c) integrable_on s"
-by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
-
-lemma integrable_on_cdivide_iff [simp]:
-  fixes f :: "_ \<Rightarrow> 'b :: real_normed_field"
-  assumes "c \<noteq> 0"
-  shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s"
-by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse)
-
-lemma has_integral_null [intro]:
-  assumes "content(cbox a b) = 0"
-  shows "(f has_integral 0) (cbox a b)"
-proof -
-  have "gauge (\<lambda>x. ball x 1)"
-    by auto
-  moreover
-  {
-    fix e :: real
-    fix p
-    assume e: "e > 0"
-    assume p: "p tagged_division_of (cbox a b)"
-    have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) = 0"
-      unfolding norm_eq_zero diff_0_right
-      using setsum_content_null[OF assms(1) p, of f] .
-    then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
-      using e by auto
-  }
-  ultimately show ?thesis
-    by (auto simp: has_integral)
-qed
-
-lemma has_integral_null_real [intro]:
-  assumes "content {a .. b::real} = 0"
-  shows "(f has_integral 0) {a .. b}"
-  by (metis assms box_real(2) has_integral_null)
-
-lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \<Longrightarrow> (f has_integral i) (cbox a b) \<longleftrightarrow> i = 0"
-  by (auto simp add: has_integral_null dest!: integral_unique)
-
-lemma integral_null [simp]: "content (cbox a b) = 0 \<Longrightarrow> integral (cbox a b) f = 0"
-  by (metis has_integral_null integral_unique)
-
-lemma integrable_on_null [intro]: "content (cbox a b) = 0 \<Longrightarrow> f integrable_on (cbox a b)"
-  by (simp add: has_integral_integrable)
-
-lemma has_integral_empty[intro]: "(f has_integral 0) {}"
-  by (simp add: has_integral_is_0)
-
-lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \<longleftrightarrow> i = 0"
-  by (auto simp add: has_integral_empty has_integral_unique)
-
-lemma integrable_on_empty[intro]: "f integrable_on {}"
-  unfolding integrable_on_def by auto
-
-lemma integral_empty[simp]: "integral {} f = 0"
-  by (rule integral_unique) (rule has_integral_empty)
-
-lemma has_integral_refl[intro]:
-  fixes a :: "'a::euclidean_space"
-  shows "(f has_integral 0) (cbox a a)"
-    and "(f has_integral 0) {a}"
-proof -
-  have *: "{a} = cbox a a"
-    apply (rule set_eqI)
-    unfolding mem_box singleton_iff euclidean_eq_iff[where 'a='a]
-    apply safe
-    prefer 3
-    apply (erule_tac x=b in ballE)
-    apply (auto simp add: field_simps)
-    done
-  show "(f has_integral 0) (cbox a a)" "(f has_integral 0) {a}"
-    unfolding *
-    apply (rule_tac[!] has_integral_null)
-    unfolding content_eq_0_interior
-    unfolding interior_cbox
-    using box_sing
-    apply auto
-    done
-qed
-
-lemma integrable_on_refl[intro]: "f integrable_on cbox a a"
-  unfolding integrable_on_def by auto
-
-lemma integral_refl [simp]: "integral (cbox a a) f = 0"
-  by (rule integral_unique) auto
-
-lemma integral_singleton [simp]: "integral {a} f = 0"
-  by auto
-
-lemma integral_blinfun_apply:
-  assumes "f integrable_on s"
-  shows "integral s (\<lambda>x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)"
-  by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def)
-
-lemma blinfun_apply_integral:
-  assumes "f integrable_on s"
-  shows "blinfun_apply (integral s f) x = integral s (\<lambda>y. blinfun_apply (f y) x)"
-  by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong)
-
-lemma has_integral_componentwise_iff:
-  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
-  shows "(f has_integral y) A \<longleftrightarrow> (\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
-proof safe
-  fix b :: 'b assume "(f has_integral y) A"
-  from has_integral_linear[OF this(1) bounded_linear_component, of b]
-    show "((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A" by (simp add: o_def)
-next
-  assume "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
-  hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral ((y \<bullet> b) *\<^sub>R b)) A"
-    by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
-  hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. (y \<bullet> b) *\<^sub>R b)) A"
-    by (intro has_integral_setsum) (simp_all add: o_def)
-  thus "(f has_integral y) A" by (simp add: euclidean_representation)
-qed
-
-lemma has_integral_componentwise:
-  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
-  shows "(\<And>b. b \<in> Basis \<Longrightarrow> ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A) \<Longrightarrow> (f has_integral y) A"
-  by (subst has_integral_componentwise_iff) blast
-
-lemma integrable_componentwise_iff:
-  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
-  shows "f integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)"
-proof
-  assume "f integrable_on A"
-  then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def)
-  hence "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)"
-    by (subst (asm) has_integral_componentwise_iff)
-  thus "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" by (auto simp: integrable_on_def)
-next
-  assume "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)"
-  then obtain y where "\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral y b) A"
-    unfolding integrable_on_def by (subst (asm) bchoice_iff) blast
-  hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral (y b *\<^sub>R b)) A"
-    by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left)
-  hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. y b *\<^sub>R b)) A"
-    by (intro has_integral_setsum) (simp_all add: o_def)
-  thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation)
-qed
-
-lemma integrable_componentwise:
-  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
-  shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) integrable_on A) \<Longrightarrow> f integrable_on A"
-  by (subst integrable_componentwise_iff) blast
-
-lemma integral_componentwise:
-  fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space"
-  assumes "f integrable_on A"
-  shows "integral A f = (\<Sum>b\<in>Basis. integral A (\<lambda>x. (f x \<bullet> b) *\<^sub>R b))"
-proof -
-  from assms have integrable: "\<forall>b\<in>Basis. (\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. (f x \<bullet> b)) integrable_on A"
-    by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI)
-       (simp_all add: bounded_linear_scaleR_left)
-  have "integral A f = integral A (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b)"
-    by (simp add: euclidean_representation)
-  also from integrable have "\<dots> = (\<Sum>a\<in>Basis. integral A (\<lambda>x. (f x \<bullet> a) *\<^sub>R a))"
-    by (subst integral_setsum) (simp_all add: o_def)
-  finally show ?thesis .
-qed
-
-lemma integrable_component:
-  "f integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (y :: 'b :: euclidean_space)) integrable_on A"
-  by (drule integrable_linear[OF _ bounded_linear_component[of y]]) (simp add: o_def)
-
-
-
-subsection \<open>Cauchy-type criterion for integrability.\<close>
-
-(* XXXXXXX *)
-lemma integrable_cauchy:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::{real_normed_vector,complete_space}"
-  shows "f integrable_on cbox a b \<longleftrightarrow>
-    (\<forall>e>0.\<exists>d. gauge d \<and>
-      (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<and> d fine p1 \<and>
-        p2 tagged_division_of (cbox a b) \<and> d fine p2 \<longrightarrow>
-        norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 -
-        setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) < e))"
-  (is "?l = (\<forall>e>0. \<exists>d. ?P e d)")
-proof
-  assume ?l
-  then guess y unfolding integrable_on_def has_integral .. note y=this
-  show "\<forall>e>0. \<exists>d. ?P e d"
-  proof (clarify, goal_cases)
-    case (1 e)
-    then have "e/2 > 0" by auto
-    then guess d
-      apply -
-      apply (drule y[rule_format])
-      apply (elim exE conjE)
-      done
-    note d=this[rule_format]
-    show ?case
-    proof (rule_tac x=d in exI, clarsimp simp: d)
-      fix p1 p2
-      assume as: "p1 tagged_division_of (cbox a b)" "d fine p1"
-                 "p2 tagged_division_of (cbox a b)" "d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-        apply (rule dist_triangle_half_l[where y=y,unfolded dist_norm])
-        using d(2)[OF conjI[OF as(1-2)]] d(2)[OF conjI[OF as(3-4)]] .
-    qed
-  qed
-next
-  assume "\<forall>e>0. \<exists>d. ?P e d"
-  then have "\<forall>n::nat. \<exists>d. ?P (inverse(of_nat (n + 1))) d"
-    by auto
-  from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format],rule_format]
-  have "\<And>n. gauge (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}})"
-    apply (rule gauge_inters)
-    using d(1)
-    apply auto
-    done
-  then have "\<forall>n. \<exists>p. p tagged_division_of (cbox a b) \<and> (\<lambda>x. \<Inter>{d i x |i. i \<in> {0..n}}) fine p"
-    by (meson fine_division_exists)
-  from choice[OF this] guess p .. note p = conjunctD2[OF this[rule_format]]
-  have dp: "\<And>i n. i\<le>n \<Longrightarrow> d i fine p n"
-    using p(2) unfolding fine_inters by auto
-  have "Cauchy (\<lambda>n. setsum (\<lambda>(x,k). content k *\<^sub>R (f x)) (p n))"
-  proof (rule CauchyI, goal_cases)
-    case (1 e)
-    then guess N unfolding real_arch_inverse[of e] .. note N=this
-    show ?case
-      apply (rule_tac x=N in exI)
-    proof clarify
-      fix m n
-      assume mn: "N \<le> m" "N \<le> n"
-      have *: "N = (N - 1) + 1" using N by auto
-      show "norm ((\<Sum>(x, k)\<in>p m. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p n. content k *\<^sub>R f x)) < e"
-        apply (rule less_trans[OF _ N[THEN conjunct2,THEN conjunct2]])
-        apply(subst *)
-        using dp p(1) mn d(2) by auto
-    qed
-  qed
-  then guess y unfolding convergent_eq_cauchy[symmetric] .. note y=this[THEN LIMSEQ_D]
-  show ?l
-    unfolding integrable_on_def has_integral
-  proof (rule_tac x=y in exI, clarify)
-    fix e :: real
-    assume "e>0"
-    then have *:"e/2 > 0" by auto
-    then guess N1 unfolding real_arch_inverse[of "e/2"] .. note N1=this
-    then have N1': "N1 = N1 - 1 + 1"
-      by auto
-    guess N2 using y[OF *] .. note N2=this
-    have "gauge (d (N1 + N2))"
-      using d by auto
-    moreover
-    {
-      fix q
-      assume as: "q tagged_division_of (cbox a b)" "d (N1 + N2) fine q"
-      have *: "inverse (of_nat (N1 + N2 + 1)) < e / 2"
-        apply (rule less_trans)
-        using N1
-        apply auto
-        done
-      have "norm ((\<Sum>(x, k)\<in>q. content k *\<^sub>R f x) - y) < e"
-        apply (rule norm_triangle_half_r)
-        apply (rule less_trans[OF _ *])
-        apply (subst N1', rule d(2)[of "p (N1+N2)"])
-        using N1' as(1) as(2) dp
-        apply (simp add: \<open>\<forall>x. p x tagged_division_of cbox a b \<and> (\<lambda>xa. \<Inter>{d i xa |i. i \<in> {0..x}}) fine p x\<close>)
-        using N2 le_add2 by blast
-    }
-    ultimately show "\<exists>d. gauge d \<and>
-      (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-        norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - y) < e)"
-      by (rule_tac x="d (N1 + N2)" in exI) auto
-  qed
-qed
-
-
-subsection \<open>Additivity of integral on abutting intervals.\<close>
-
-lemma tagged_division_split_left_inj:
-  fixes x1 :: "'a::euclidean_space"
-  assumes d: "d tagged_division_of i"
-    and k12: "(x1, k1) \<in> d"
-             "(x2, k2) \<in> d"
-             "k1 \<noteq> k2"
-             "k1 \<inter> {x. x\<bullet>k \<le> c} = k2 \<inter> {x. x\<bullet>k \<le> c}"
-             "k \<in> Basis"
-  shows "content (k1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
-proof -
-  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
-    by force
-  show ?thesis
-    using k12
-    by (fastforce intro!:  division_split_left_inj[OF division_of_tagged_division[OF d]] *)
-qed
-
-lemma tagged_division_split_right_inj:
-  fixes x1 :: "'a::euclidean_space"
-  assumes d: "d tagged_division_of i"
-    and k12: "(x1, k1) \<in> d"
-             "(x2, k2) \<in> d"
-             "k1 \<noteq> k2"
-             "k1 \<inter> {x. x\<bullet>k \<ge> c} = k2 \<inter> {x. x\<bullet>k \<ge> c}"
-             "k \<in> Basis"
-  shows "content (k1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
-proof -
-  have *: "\<And>a b c. (a,b) \<in> c \<Longrightarrow> b \<in> snd ` c"
-    by force
-  show ?thesis
-    using k12
-    by (fastforce intro!:  division_split_right_inj[OF division_of_tagged_division[OF d]] *)
-qed
-
-lemma has_integral_split:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes fi: "(f has_integral i) (cbox a b \<inter> {x. x\<bullet>k \<le> c})"
-      and fj: "(f has_integral j) (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-      and k: "k \<in> Basis"
-  shows "(f has_integral (i + j)) (cbox a b)"
-proof (unfold has_integral, rule, rule, goal_cases)
-  case (1 e)
-  then have e: "e/2 > 0"
-    by auto
-    obtain d1
-    where d1: "gauge d1"
-      and d1norm:
-        "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c};
-               d1 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. content k *\<^sub>R f x) - i) < e / 2"
-       apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e])
-       apply (simp add: interval_split[symmetric] k)
-       done
-    obtain d2
-    where d2: "gauge d2"
-      and d2norm:
-        "\<And>p. \<lbrakk>p tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k};
-               d2 fine p\<rbrakk> \<Longrightarrow> norm ((\<Sum>(x, k) \<in> p. content k *\<^sub>R f x) - j) < e / 2"
-       apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e])
-       apply (simp add: interval_split[symmetric] k)
-       done
-  let ?d = "\<lambda>x. if x\<bullet>k = c then (d1 x \<inter> d2 x) else ball x \<bar>x\<bullet>k - c\<bar> \<inter> d1 x \<inter> d2 x"
-  have "gauge ?d"
-    using d1 d2 unfolding gauge_def by auto
-  then show ?case
-  proof (rule_tac x="?d" in exI, safe)
-    fix p
-    assume "p tagged_division_of (cbox a b)" "?d fine p"
-    note p = this tagged_division_ofD[OF this(1)]
-    have xk_le_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<le> c"
-    proof -
-      fix x kk
-      assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}"
-      show "x\<bullet>k \<le> c"
-      proof (rule ccontr)
-        assume **: "\<not> ?thesis"
-        from this[unfolded not_le]
-        have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
-          using p(2)[unfolded fine_def, rule_format,OF as] by auto
-        with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<le> c"
-          by blast
-        then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>"
-          using Basis_le_norm[OF k, of "x - y"]
-          by (auto simp add: dist_norm inner_diff_left intro: le_less_trans)
-        with y show False
-          using ** by (auto simp add: field_simps)
-      qed
-    qed
-    have xk_ge_c: "\<And>x kk. (x, kk) \<in> p \<Longrightarrow> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {} \<Longrightarrow> x\<bullet>k \<ge> c"
-    proof -
-      fix x kk
-      assume as: "(x, kk) \<in> p" and kk: "kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}"
-      show "x\<bullet>k \<ge> c"
-      proof (rule ccontr)
-        assume **: "\<not> ?thesis"
-        from this[unfolded not_le] have "kk \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>"
-          using p(2)[unfolded fine_def,rule_format,OF as,unfolded split_conv] by auto
-        with kk obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<ge> c"
-          by blast
-        then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>"
-          using Basis_le_norm[OF k, of "x - y"]
-          by (auto simp add: dist_norm inner_diff_left intro: le_less_trans)
-        with y show False
-          using ** by (auto simp add: field_simps)
-      qed
-    qed
-
-    have lem1: "\<And>f P Q. (\<forall>x k. (x, k) \<in> {(x, f k) | x k. P x k} \<longrightarrow> Q x k) \<longleftrightarrow>
-                         (\<forall>x k. P x k \<longrightarrow> Q x (f k))"
-      by auto
-    have fin_finite: "finite {(x,f k) | x k. (x,k) \<in> s \<and> P x k}" if "finite s" for f s P
-    proof -
-      from that have "finite ((\<lambda>(x, k). (x, f k)) ` s)"
-        by auto
-      then show ?thesis
-        by (rule rev_finite_subset) auto
-    qed
-    { fix g :: "'a set \<Rightarrow> 'a set"
-      fix i :: "'a \<times> 'a set"
-      assume "i \<in> (\<lambda>(x, k). (x, g k)) ` p - {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
-      then obtain x k where xk:
-              "i = (x, g k)"  "(x, k) \<in> p"
-              "(x, g k) \<notin> {(x, g k) |x k. (x, k) \<in> p \<and> g k \<noteq> {}}"
-          by auto
-      have "content (g k) = 0"
-        using xk using content_empty by auto
-      then have "(\<lambda>(x, k). content k *\<^sub>R f x) i = 0"
-        unfolding xk split_conv by auto
-    } note [simp] = this
-    have lem3: "\<And>g :: 'a set \<Rightarrow> 'a set. finite p \<Longrightarrow>
-                  setsum (\<lambda>(x, k). content k *\<^sub>R f x) {(x,g k) |x k. (x,k) \<in> p \<and> g k \<noteq> {}} =
-                  setsum (\<lambda>(x, k). content k *\<^sub>R f x) ((\<lambda>(x, k). (x, g k)) ` p)"
-      by (rule setsum.mono_neutral_left) auto
-    let ?M1 = "{(x, kk \<inter> {x. x\<bullet>k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
-    have d1_fine: "d1 fine ?M1"
-      by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm)
-    have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2"
-    proof (rule d1norm [OF tagged_division_ofI d1_fine])
-      show "finite ?M1"
-        by (rule fin_finite p(3))+
-      show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = cbox a b \<inter> {x. x\<bullet>k \<le> c}"
-        unfolding p(8)[symmetric] by auto
-      fix x l
-      assume xl: "(x, l) \<in> ?M1"
-      then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this
-      show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<le> c}"
-        unfolding xl'
-        using p(4-6)[OF xl'(3)] using xl'(4)
-        using xk_le_c[OF xl'(3-4)] by auto
-      show "\<exists>a b. l = cbox a b"
-        unfolding xl'
-        using p(6)[OF xl'(3)]
-        by (fastforce simp add: interval_split[OF k,where c=c])
-      fix y r
-      let ?goal = "interior l \<inter> interior r = {}"
-      assume yr: "(y, r) \<in> ?M1"
-      then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this
-      assume as: "(x, l) \<noteq> (y, r)"
-      show "interior l \<inter> interior r = {}"
-      proof (cases "l' = r' \<longrightarrow> x' = y'")
-        case False
-        then show ?thesis
-          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      next
-        case True
-        then have "l' \<noteq> r'"
-          using as unfolding xl' yr' by auto
-        then show ?thesis
-          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      qed
-    qed
-    moreover
-    let ?M2 = "{(x,kk \<inter> {x. x\<bullet>k \<ge> c}) |x kk. (x,kk) \<in> p \<and> kk \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
-    have d2_fine: "d2 fine ?M2"
-      by (force intro: fineI dest: fineD[OF p(2)] simp add: split: if_split_asm)
-    have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2"
-    proof (rule d2norm [OF tagged_division_ofI d2_fine])
-      show "finite ?M2"
-        by (rule fin_finite p(3))+
-      show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = cbox a b \<inter> {x. x\<bullet>k \<ge> c}"
-        unfolding p(8)[symmetric] by auto
-      fix x l
-      assume xl: "(x, l) \<in> ?M2"
-      then guess x' l' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note xl'=this
-      show "x \<in> l" "l \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
-        unfolding xl'
-        using p(4-6)[OF xl'(3)] xl'(4) xk_ge_c[OF xl'(3-4)]
-        by auto
-      show "\<exists>a b. l = cbox a b"
-        unfolding xl'
-        using p(6)[OF xl'(3)]
-        by (fastforce simp add: interval_split[OF k, where c=c])
-      fix y r
-      let ?goal = "interior l \<inter> interior r = {}"
-      assume yr: "(y, r) \<in> ?M2"
-      then guess y' r' unfolding mem_Collect_eq unfolding prod.inject by (elim exE conjE) note yr'=this
-      assume as: "(x, l) \<noteq> (y, r)"
-      show "interior l \<inter> interior r = {}"
-      proof (cases "l' = r' \<longrightarrow> x' = y'")
-        case False
-        then show ?thesis
-          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      next
-        case True
-        then have "l' \<noteq> r'"
-          using as unfolding xl' yr' by auto
-        then show ?thesis
-          using p(7)[OF xl'(3) yr'(3)] using as unfolding xl' yr' by auto
-      qed
-    qed
-    ultimately
-    have "norm (((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j)) < e/2 + e/2"
-      using norm_add_less by blast
-    also {
-      have eq0: "\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0"
-        using scaleR_zero_left by auto
-      have cont_eq: "\<And>g. (\<lambda>(x,l). content l *\<^sub>R f x) \<circ> (\<lambda>(x,l). (x,g l)) = (\<lambda>(x,l). content (g l) *\<^sub>R f x)"
-        by auto
-      have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) =
-        (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)"
-        by auto
-      also have "\<dots> = (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
-        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
-        unfolding lem3[OF p(3)]
-        by (subst setsum.reindex_nontrivial[OF p(3)], auto intro!: k eq0 tagged_division_split_left_inj[OF p(1)] tagged_division_split_right_inj[OF p(1)]
-              simp: cont_eq)+
-      also note setsum.distrib[symmetric]
-      also have "\<And>x. x \<in> p \<Longrightarrow>
-                    (\<lambda>(x,ka). content (ka \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) x +
-                    (\<lambda>(x,ka). content (ka \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) x =
-                    (\<lambda>(x,ka). content ka *\<^sub>R f x) x"
-      proof clarify
-        fix a b
-        assume "(a, b) \<in> p"
-        from p(6)[OF this] guess u v by (elim exE) note uv=this
-        then show "content (b \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f a + content (b \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f a =
-          content b *\<^sub>R f a"
-          unfolding scaleR_left_distrib[symmetric]
-          unfolding uv content_split[OF k,of u v c]
-          by auto
-      qed
-      note setsum.cong [OF _ this]
-      finally have "(\<Sum>(x, k)\<in>{(x, kk \<inter> {x. x \<bullet> k \<le> c}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}}. content k *\<^sub>R f x) - i +
-        ((\<Sum>(x, k)\<in>{(x, kk \<inter> {x. c \<le> x \<bullet> k}) |x kk. (x, kk) \<in> p \<and> kk \<inter> {x. c \<le> x \<bullet> k} \<noteq> {}}. content k *\<^sub>R f x) - j) =
-        (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)"
-        by auto
-    }
-    finally show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e"
-      by auto
-  qed
-qed
-
-
-subsection \<open>A sort of converse, integrability on subintervals.\<close>
-
-lemma tagged_division_union_interval:
-  fixes a :: "'a::euclidean_space"
-  assumes "p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> (c::real)})"
-    and "p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-    and k: "k \<in> Basis"
-  shows "(p1 \<union> p2) tagged_division_of (cbox a b)"
-proof -
-  have *: "cbox a b = (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<union> (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
-    by auto
-  show ?thesis
-    apply (subst *)
-    apply (rule tagged_division_union[OF assms(1-2)])
-    unfolding interval_split[OF k] interior_cbox
-    using k
-    apply (auto simp add: box_def elim!: ballE[where x=k])
-    done
-qed
-
-lemma tagged_division_union_interval_real:
-  fixes a :: real
-  assumes "p1 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<le> (c::real)})"
-    and "p2 tagged_division_of ({a .. b} \<inter> {x. x\<bullet>k \<ge> c})"
-    and k: "k \<in> Basis"
-  shows "(p1 \<union> p2) tagged_division_of {a .. b}"
-  using assms
-  unfolding box_real[symmetric]
-  by (rule tagged_division_union_interval)
-
-lemma has_integral_separate_sides:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "(f has_integral i) (cbox a b)"
-    and "e > 0"
-    and k: "k \<in> Basis"
-  obtains d where "gauge d"
-    "\<forall>p1 p2. p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and>
-        p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2 \<longrightarrow>
-        norm ((setsum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + setsum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e"
-proof -
-  guess d using has_integralD[OF assms(1-2)] . note d=this
-  { fix p1 p2
-    assume "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1"
-    note p1=tagged_division_ofD[OF this(1)] this
-    assume "p2 tagged_division_of (cbox a b) \<inter> {x. c \<le> x \<bullet> k}" "d fine p2"
-    note p2=tagged_division_ofD[OF this(1)] this
-    note tagged_division_union_interval[OF p1(7) p2(7)] note p12 = tagged_division_ofD[OF this] this
-    { fix a b
-      assume ab: "(a, b) \<in> p1 \<inter> p2"
-      have "(a, b) \<in> p1"
-        using ab by auto
-      with p1 obtain u v where uv: "b = cbox u v" by auto
-      have "b \<subseteq> {x. x\<bullet>k = c}"
-        using ab p1(3)[of a b] p2(3)[of a b] by fastforce
-      moreover
-      have "interior {x::'a. x \<bullet> k = c} = {}"
-      proof (rule ccontr)
-        assume "\<not> ?thesis"
-        then obtain x where x: "x \<in> interior {x::'a. x\<bullet>k = c}"
-          by auto
-        then guess e unfolding mem_interior .. note e=this
-        have x: "x\<bullet>k = c"
-          using x interior_subset by fastforce
-        have *: "\<And>i. i \<in> Basis \<Longrightarrow> \<bar>(x - (x + (e / 2) *\<^sub>R k)) \<bullet> i\<bar> = (if i = k then e/2 else 0)"
-          using e k by (auto simp: inner_simps inner_not_same_Basis)
-        have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (e / 2 ) *\<^sub>R k)) \<bullet> i\<bar>) =
-              (\<Sum>i\<in>Basis. (if i = k then e / 2 else 0))"
-          using "*" by (blast intro: setsum.cong)
-        also have "\<dots> < e"
-          apply (subst setsum.delta)
-          using e
-          apply auto
-          done
-        finally have "x + (e/2) *\<^sub>R k \<in> ball x e"
-          unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1])
-        then have "x + (e/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}"
-          using e by auto
-        then show False
-          unfolding mem_Collect_eq using e x k by (auto simp: inner_simps)
-      qed
-      ultimately have "content b = 0"
-        unfolding uv content_eq_0_interior
-        using interior_mono by blast
-      then have "content b *\<^sub>R f a = 0"
-        by auto
-    }
-    then have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) =
-               norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)"
-      by (subst setsum.union_inter_neutral) (auto simp: p1 p2)
-    also have "\<dots> < e"
-      by (rule k d(2) p12 fine_union p1 p2)+
-    finally have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" .
-   }
-  then show ?thesis
-    by (auto intro: that[of d] d elim: )
-qed
-
-lemma integrable_split[intro]:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
-  assumes "f integrable_on cbox a b"
-    and k: "k \<in> Basis"
-  shows "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<le> c})" (is ?t1)
-    and "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})" (is ?t2)
-proof -
-  guess y using assms(1) unfolding integrable_on_def .. note y=this
-  define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)"
-  define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i)"
-  show ?t1 ?t2
-    unfolding interval_split[OF k] integrable_cauchy
-    unfolding interval_split[symmetric,OF k]
-  proof (rule_tac[!] allI impI)+
-    fix e :: real
-    assume "e > 0"
-    then have "e/2>0"
-      by auto
-    from has_integral_separate_sides[OF y this k,of c] guess d . note d=this[rule_format]
-    let ?P = "\<lambda>A. \<exists>d. gauge d \<and> (\<forall>p1 p2. p1 tagged_division_of (cbox a b) \<inter> A \<and> d fine p1 \<and>
-      p2 tagged_division_of (cbox a b) \<inter> A \<and> d fine p2 \<longrightarrow>
-      norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e)"
-    show "?P {x. x \<bullet> k \<le> c}"
-    proof (rule_tac x=d in exI, clarsimp simp add: d)
-      fix p1 p2
-      assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1"
-                 "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-      proof (rule fine_division_exists[OF d(1), of a' b] )
-        fix p
-        assume "p tagged_division_of cbox a' b" "d fine p"
-        then show ?thesis
-          using as norm_triangle_half_l[OF d(2)[of p1 p] d(2)[of p2 p]]
-          unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
-          by (auto simp add: algebra_simps)
-      qed
-    qed
-    show "?P {x. x \<bullet> k \<ge> c}"
-    proof (rule_tac x=d in exI, clarsimp simp add: d)
-      fix p1 p2
-      assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p1"
-                 "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p2"
-      show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e"
-      proof (rule fine_division_exists[OF d(1), of a b'] )
-        fix p
-        assume "p tagged_division_of cbox a b'" "d fine p"
-        then show ?thesis
-          using as norm_triangle_half_l[OF d(2)[of p p1] d(2)[of p p2]]
-          unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric]
-          by (auto simp add: algebra_simps)
-      qed
-    qed
-  qed
-qed
-
-lemma operative_integral:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  shows "comm_monoid.operative (lift_option op +) (Some 0) (\<lambda>i. if f integrable_on i then Some (integral i f) else None)"
-  unfolding comm_monoid.operative_def[OF add.comm_monoid_lift_option]
-proof safe
-  fix a b c
-  fix k :: 'a
-  assume k: "k \<in> Basis"
-  show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) =
-        lift_option op + (if f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c} then Some (integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f) else None)
-        (if f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k} then Some (integral (cbox a b \<inter> {x. c \<le> x \<bullet> k}) f) else None)"
-  proof (cases "f integrable_on cbox a b")
-    case True
-    with k show ?thesis
-      apply (simp add: integrable_split)
-      apply (rule integral_unique [OF has_integral_split[OF _ _ k]])
-      apply (auto intro: integrable_integral)
-      done
-  next
-    case False
-    have "\<not> (f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<or> \<not> ( f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k})"
-    proof (rule ccontr)
-      assume "\<not> ?thesis"
-      then have "f integrable_on cbox a b"
-        unfolding integrable_on_def
-        apply (rule_tac x="integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f + integral (cbox a b \<inter> {x. x \<bullet> k \<ge> c}) f" in exI)
-        apply (rule has_integral_split[OF _ _ k])
-        apply (auto intro: integrable_integral)
-        done
-      then show False
-        using False by auto
-    qed
-    then show ?thesis
-      using False by auto
-  qed
-next
-  fix a b :: 'a
-  assume "content (cbox a b) = 0"
-  then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0"
-    using has_integral_null_eq
-    by (auto simp: integrable_on_null)
-qed
-
-subsection \<open>Finally, the integral of a constant\<close>
-
-lemma has_integral_const [intro]:
-  fixes a b :: "'a::euclidean_space"
-  shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)"
-  apply (auto intro!: exI [where x="\<lambda>x. ball x 1"] simp: split_def has_integral)
-  apply (subst scaleR_left.setsum[symmetric, unfolded o_def])
-  apply (subst additive_content_tagged_division[unfolded split_def])
-  apply auto
-  done
-
-lemma has_integral_const_real [intro]:
-  fixes a b :: real
-  shows "((\<lambda>x. c) has_integral (content {a .. b} *\<^sub>R c)) {a .. b}"
-  by (metis box_real(2) has_integral_const)
-
-lemma integral_const [simp]:
-  fixes a b :: "'a::euclidean_space"
-  shows "integral (cbox a b) (\<lambda>x. c) = content (cbox a b) *\<^sub>R c"
-  by (rule integral_unique) (rule has_integral_const)
-
-lemma integral_const_real [simp]:
-  fixes a b :: real
-  shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
-  by (metis box_real(2) integral_const)
-
-
-subsection \<open>Bounds on the norm of Riemann sums and the integral itself.\<close>
-
-lemma dsum_bound:
-  assumes "p division_of (cbox a b)"
-    and "norm c \<le> e"
-  shows "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)"
-proof -
-  have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = setsum content p"
-    apply (rule setsum.cong)
-    using assms
-    apply simp
-    apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4))
-    done
-  have e: "0 \<le> e"
-    using assms(2) norm_ge_zero order_trans by blast
-  have "norm (setsum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))"
-    using norm_setsum by blast
-  also have "...  \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)"
-    apply (simp add: setsum_right_distrib[symmetric] mult.commute)
-    using assms(2) mult_right_mono by blast
-  also have "... \<le> e * content (cbox a b)"
-    apply (rule mult_left_mono [OF _ e])
-    apply (simp add: sumeq)
-    using additive_content_division assms(1) eq_iff apply blast
-    done
-  finally show ?thesis .
-qed
-
-lemma rsum_bound:
-  assumes p: "p tagged_division_of (cbox a b)"
-      and "\<forall>x\<in>cbox a b. norm (f x) \<le> e"
-    shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content (cbox a b)"
-proof (cases "cbox a b = {}")
-  case True show ?thesis
-    using p unfolding True tagged_division_of_trivial by auto
-next
-  case False
-  then have e: "e \<ge> 0"
-    by (meson ex_in_conv assms(2) norm_ge_zero order_trans)
-  have setsum_le: "setsum (content \<circ> snd) p \<le> content (cbox a b)"
-    unfolding additive_content_tagged_division[OF p, symmetric] split_def
-    by (auto intro: eq_refl)
-  have con: "\<And>xk. xk \<in> p \<Longrightarrow> 0 \<le> content (snd xk)"
-    using tagged_division_ofD(4) [OF p] content_pos_le
-    by force
-  have norm: "\<And>xk. xk \<in> p \<Longrightarrow> norm (f (fst xk)) \<le> e"
-    unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms
-    by (metis prod.collapse subset_eq)
-  have "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))"
-    by (rule norm_setsum)
-  also have "...  \<le> e * content (cbox a b)"
-    unfolding split_def norm_scaleR
-    apply (rule order_trans[OF setsum_mono])
-    apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e])
-    apply (metis norm)
-    unfolding setsum_left_distrib[symmetric]
-    using con setsum_le
-    apply (auto simp: mult.commute intro: mult_left_mono [OF _ e])
-    done
-  finally show ?thesis .
-qed
-
-lemma rsum_diff_bound:
-  assumes "p tagged_division_of (cbox a b)"
-    and "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e"
-  shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - setsum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le>
-         e * content (cbox a b)"
-  apply (rule order_trans[OF _ rsum_bound[OF assms]])
-  apply (simp add: split_def scaleR_diff_right setsum_subtractf eq_refl)
-  done
-
-lemma has_integral_bound:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "0 \<le> B"
-      and "(f has_integral i) (cbox a b)"
-      and "\<forall>x\<in>cbox a b. norm (f x) \<le> B"
-    shows "norm i \<le> B * content (cbox a b)"
-proof (rule ccontr)
-  assume "\<not> ?thesis"
-  then have *: "norm i - B * content (cbox a b) > 0"
-    by auto
-  from assms(2)[unfolded has_integral,rule_format,OF *]
-  guess d by (elim exE conjE) note d=this[rule_format]
-  from fine_division_exists[OF this(1), of a b] guess p . note p=this
-  have *: "\<And>s B. norm s \<le> B \<Longrightarrow> \<not> norm (s - i) < norm i - B"
-    unfolding not_less
-    by (metis norm_triangle_sub[of i] add.commute le_less_trans less_diff_eq linorder_not_le norm_minus_commute)
-  show False
-    using d(2)[OF conjI[OF p]] *[OF rsum_bound[OF p(1) assms(3)]] by auto
-qed
-
-corollary has_integral_bound_real:
-  fixes f :: "real \<Rightarrow> 'b::real_normed_vector"
-  assumes "0 \<le> B"
-      and "(f has_integral i) {a .. b}"
-      and "\<forall>x\<in>{a .. b}. norm (f x) \<le> B"
-    shows "norm i \<le> B * content {a .. b}"
-  by (metis assms box_real(2) has_integral_bound)
-
-corollary integrable_bound:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "0 \<le> B"
-      and "f integrable_on (cbox a b)"
-      and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B"
-    shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)"
-by (metis integrable_integral has_integral_bound assms)
-
-
-subsection \<open>Similar theorems about relationship among components.\<close>
-
-lemma rsum_component_le:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "p tagged_division_of (cbox a b)"
-      and "\<forall>x\<in>cbox a b. (f x)\<bullet>i \<le> (g x)\<bullet>i"
-    shows "(setsum (\<lambda>(x,k). content k *\<^sub>R f x) p)\<bullet>i \<le> (setsum (\<lambda>(x,k). content k *\<^sub>R g x) p)\<bullet>i"
-unfolding inner_setsum_left
-proof (rule setsum_mono, clarify)
-  fix a b
-  assume ab: "(a, b) \<in> p"
-  note tagged = tagged_division_ofD(2-4)[OF assms(1) ab]
-  from this(3) guess u v by (elim exE) note b=this
-  show "(content b *\<^sub>R f a) \<bullet> i \<le> (content b *\<^sub>R g a) \<bullet> i"
-    unfolding b inner_simps real_scaleR_def
-    apply (rule mult_left_mono)
-    using assms(2) tagged
-    by (auto simp add: content_pos_le)
-qed
-
-lemma has_integral_component_le:
-  fixes f g :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes k: "k \<in> Basis"
-  assumes "(f has_integral i) s" "(g has_integral j) s"
-    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
-  shows "i\<bullet>k \<le> j\<bullet>k"
-proof -
-  have lem: "i\<bullet>k \<le> j\<bullet>k"
-    if f_i: "(f has_integral i) (cbox a b)"
-    and g_j: "(g has_integral j) (cbox a b)"
-    and le: "\<forall>x\<in>cbox a b. (f x)\<bullet>k \<le> (g x)\<bullet>k"
-    for a b i and j :: 'b and f g :: "'a \<Rightarrow> 'b"
-  proof (rule ccontr)
-    assume "\<not> ?thesis"
-    then have *: "0 < (i\<bullet>k - j\<bullet>k) / 3"
-      by auto
-    guess d1 using f_i[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d1=this[rule_format]
-    guess d2 using g_j[unfolded has_integral,rule_format,OF *] by (elim exE conjE) note d2=this[rule_format]
-    obtain p where p: "p tagged_division_of cbox a b" "d1 fine p" "d2 fine p"
-       using fine_division_exists[OF gauge_inter[OF d1(1) d2(1)], of a b] unfolding fine_inter
-       by metis
-    note le_less_trans[OF Basis_le_norm[OF k]]
-    then have "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
-              "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3"
-      using  k norm_bound_Basis_lt d1 d2 p
-      by blast+
-    then show False
-      unfolding inner_simps
-      using rsum_component_le[OF p(1) le]
-      by (simp add: abs_real_def split: if_split_asm)
-  qed
-  show ?thesis
-  proof (cases "\<exists>a b. s = cbox a b")
-    case True
-    with lem assms show ?thesis
-      by auto
-  next
-    case False
-    show ?thesis
-    proof (rule ccontr)
-      assume "\<not> i\<bullet>k \<le> j\<bullet>k"
-      then have ij: "(i\<bullet>k - j\<bullet>k) / 3 > 0"
-        by auto
-      note has_integral_altD[OF _ False this]
-      from this[OF assms(2)] this[OF assms(3)] guess B1 B2 . note B=this[rule_format]
-      have "bounded (ball 0 B1 \<union> ball (0::'a) B2)"
-        unfolding bounded_Un by(rule conjI bounded_ball)+
-      from bounded_subset_cbox[OF this] guess a b by (elim exE)
-      note ab = conjunctD2[OF this[unfolded Un_subset_iff]]
-      guess w1 using B(2)[OF ab(1)] .. note w1=conjunctD2[OF this]
-      guess w2 using B(4)[OF ab(2)] .. note w2=conjunctD2[OF this]
-      have *: "\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False"
-        by (simp add: abs_real_def split: if_split_asm)
-      note le_less_trans[OF Basis_le_norm[OF k]]
-      note this[OF w1(2)] this[OF w2(2)]
-      moreover
-      have "w1\<bullet>k \<le> w2\<bullet>k"
-        by (rule lem[OF w1(1) w2(1)]) (simp add: assms(4))
-      ultimately show False
-        unfolding inner_simps by(rule *)
-    qed
-  qed
-qed
-
-lemma integral_component_le:
-  fixes g f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "k \<in> Basis"
-    and "f integrable_on s" "g integrable_on s"
-    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> (g x)\<bullet>k"
-  shows "(integral s f)\<bullet>k \<le> (integral s g)\<bullet>k"
-  apply (rule has_integral_component_le)
-  using integrable_integral assms
-  apply auto
-  done
-
-lemma has_integral_component_nonneg:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "k \<in> Basis"
-    and "(f has_integral i) s"
-    and "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k"
-  shows "0 \<le> i\<bullet>k"
-  using has_integral_component_le[OF assms(1) has_integral_0 assms(2)]
-  using assms(3-)
-  by auto
-
-lemma integral_component_nonneg:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "k \<in> Basis"
-    and  "\<forall>x\<in>s. 0 \<le> (f x)\<bullet>k"
-  shows "0 \<le> (integral s f)\<bullet>k"
-proof (cases "f integrable_on s")
-  case True show ?thesis
-    apply (rule has_integral_component_nonneg)
-    using assms True
-    apply auto
-    done
-next
-  case False then show ?thesis by (simp add: not_integrable_integral)
-qed
-
-lemma has_integral_component_neg:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "k \<in> Basis"
-    and "(f has_integral i) s"
-    and "\<forall>x\<in>s. (f x)\<bullet>k \<le> 0"
-  shows "i\<bullet>k \<le> 0"
-  using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-)
-  by auto
-
-lemma has_integral_component_lbound:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "(f has_integral i) (cbox a b)"
-    and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k"
-    and "k \<in> Basis"
-  shows "B * content (cbox a b) \<le> i\<bullet>k"
-  using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\<Sum>i\<in>Basis. B *\<^sub>R i)::'b"] assms(2-)
-  by (auto simp add: field_simps)
-
-lemma has_integral_component_ubound:
-  fixes f::"'a::euclidean_space => 'b::euclidean_space"
-  assumes "(f has_integral i) (cbox a b)"
-    and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B"
-    and "k \<in> Basis"
-  shows "i\<bullet>k \<le> B * content (cbox a b)"
-  using has_integral_component_le[OF assms(3,1) has_integral_const, of "\<Sum>i\<in>Basis. B *\<^sub>R i"] assms(2-)
-  by (auto simp add: field_simps)
-
-lemma integral_component_lbound:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "f integrable_on cbox a b"
-    and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k"
-    and "k \<in> Basis"
-  shows "B * content (cbox a b) \<le> (integral(cbox a b) f)\<bullet>k"
-  apply (rule has_integral_component_lbound)
-  using assms
-  unfolding has_integral_integral
-  apply auto
-  done
-
-lemma integral_component_lbound_real:
-  assumes "f integrable_on {a ::real .. b}"
-    and "\<forall>x\<in>{a .. b}. B \<le> f(x)\<bullet>k"
-    and "k \<in> Basis"
-  shows "B * content {a .. b} \<le> (integral {a .. b} f)\<bullet>k"
-  using assms
-  by (metis box_real(2) integral_component_lbound)
-
-lemma integral_component_ubound:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes "f integrable_on cbox a b"
-    and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B"
-    and "k \<in> Basis"
-  shows "(integral (cbox a b) f)\<bullet>k \<le> B * content (cbox a b)"
-  apply (rule has_integral_component_ubound)
-  using assms
-  unfolding has_integral_integral
-  apply auto
-  done
-
-lemma integral_component_ubound_real:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "f integrable_on {a .. b}"
-    and "\<forall>x\<in>{a .. b}. f x\<bullet>k \<le> B"
-    and "k \<in> Basis"
-  shows "(integral {a .. b} f)\<bullet>k \<le> B * content {a .. b}"
-  using assms
-  by (metis box_real(2) integral_component_ubound)
-
-subsection \<open>Uniform limit of integrable functions is integrable.\<close>
-
-lemma real_arch_invD:
-  "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)"
-  by (subst(asm) real_arch_inverse)
-
-lemma integrable_uniform_limit:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "\<forall>e>0. \<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
-  shows "f integrable_on cbox a b"
-proof (cases "content (cbox a b) > 0")
-  case False then show ?thesis
-      using has_integral_null
-      by (simp add: content_lt_nz integrable_on_def)
-next
-  case True
-  have *: "\<And>P. \<forall>e>(0::real). P e \<Longrightarrow> \<forall>n::nat. P (inverse (real n + 1))"
-    by auto
-  from choice[OF *[OF assms]] guess g .. note g=conjunctD2[OF this[rule_format],rule_format]
-  from choice[OF allI[OF g(2)[unfolded integrable_on_def], of "\<lambda>x. x"]]
-  obtain i where i: "\<And>x. (g x has_integral i x) (cbox a b)"
-      by auto
-  have "Cauchy i"
-    unfolding Cauchy_def
-  proof clarify
-    fix e :: real
-    assume "e>0"
-    then have "e / 4 / content (cbox a b) > 0"
-      using True by (auto simp add: field_simps)
-    then obtain M :: nat
-         where M: "M \<noteq> 0" "0 < inverse (real_of_nat M)" "inverse (of_nat M) < e / 4 / content (cbox a b)"
-      by (subst (asm) real_arch_inverse) auto
-    show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (i m) (i n) < e"
-    proof (rule exI [where x=M], clarify)
-      fix m n
-      assume m: "M \<le> m" and n: "M \<le> n"
-      have "e/4>0" using \<open>e>0\<close> by auto
-      note * = i[unfolded has_integral,rule_format,OF this]
-      from *[of m] guess gm by (elim conjE exE) note gm=this[rule_format]
-      from *[of n] guess gn by (elim conjE exE) note gn=this[rule_format]
-      from fine_division_exists[OF gauge_inter[OF gm(1) gn(1)], of a b]
-      obtain p where p: "p tagged_division_of cbox a b" "(\<lambda>x. gm x \<inter> gn x) fine p"
-        by auto
-      { fix s1 s2 i1 and i2::'b
-        assume no: "norm(s2 - s1) \<le> e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4"
-        have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)"
-          using norm_triangle_ineq[of "i1 - s1" "s1 - i2"]
-          using norm_triangle_ineq[of "s1 - s2" "s2 - i2"]
-          by (auto simp add: algebra_simps)
-        also have "\<dots> < e"
-          using no
-          unfolding norm_minus_commute
-          by (auto simp add: algebra_simps)
-        finally have "norm (i1 - i2) < e" .
-      } note triangle3 = this
-      have finep: "gm fine p" "gn fine p"
-        using fine_inter p  by auto
-      { fix x
-        assume x: "x \<in> cbox a b"
-        have "norm (f x - g n x) + norm (f x - g m x) \<le> inverse (real n + 1) + inverse (real m + 1)"
-          using g(1)[OF x, of n] g(1)[OF x, of m] by auto
-        also have "\<dots> \<le> inverse (real M) + inverse (real M)"
-          apply (rule add_mono)
-          using M(2) m n by auto
-        also have "\<dots> = 2 / real M"
-          unfolding divide_inverse by auto
-        finally have "norm (g n x - g m x) \<le> 2 / real M"
-          using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"]
-          by (auto simp add: algebra_simps simp add: norm_minus_commute)
-      } note norm_le = this
-      have le_e2: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g n x) - (\<Sum>(x, k)\<in>p. content k *\<^sub>R g m x)) \<le> e / 2"
-        apply (rule order_trans [OF rsum_diff_bound[OF p(1), where e="2 / real M"]])
-        apply (blast intro: norm_le)
-        using M True
-        by (auto simp add: field_simps)
-      then show "dist (i m) (i n) < e"
-        unfolding dist_norm
-        using gm gn p finep
-        by (auto intro!: triangle3)
-    qed
-  qed
-  then obtain s where s: "i \<longlonglongrightarrow> s"
-    using convergent_eq_cauchy[symmetric] by blast
-  show ?thesis
-    unfolding integrable_on_def has_integral
-  proof (rule_tac x=s in exI, clarify)
-    fix e::real
-    assume e: "0 < e"
-    then have *: "e/3 > 0" by auto
-    then obtain N1 where N1: "\<forall>n\<ge>N1. norm (i n - s) < e / 3"
-      using LIMSEQ_D [OF s] by metis
-    from e True have "e / 3 / content (cbox a b) > 0"
-      by (auto simp add: field_simps)
-    from real_arch_invD[OF this] guess N2 by (elim exE conjE) note N2=this
-    from i[of "N1 + N2",unfolded has_integral,rule_format,OF *] guess g' .. note g'=conjunctD2[OF this,rule_format]
-    { fix sf sg i
-      assume no: "norm (sf - sg) \<le> e / 3"
-                 "norm(i - s) < e / 3"
-                 "norm (sg - i) < e / 3"
-      have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - i) + norm (i - s)"
-        using norm_triangle_ineq[of "sf - sg" "sg - s"]
-        using norm_triangle_ineq[of "sg -  i" " i - s"]
-        by (auto simp add: algebra_simps)
-      also have "\<dots> < e"
-        using no
-        unfolding norm_minus_commute
-        by (auto simp add: algebra_simps)
-      finally have "norm (sf - s) < e" .
-    } note lem = this
-    { fix p
-      assume p: "p tagged_division_of (cbox a b) \<and> g' fine p"
-      then have norm_less: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g (N1 + N2) x) - i (N1 + N2)) < e / 3"
-        using g' by blast
-      have "content (cbox a b) < e / 3 * (of_nat N2)"
-        using N2 unfolding inverse_eq_divide using True by (auto simp add: field_simps)
-      moreover have "e / 3 * of_nat N2 \<le> e / 3 * (of_nat (N1 + N2) + 1)"
-        using \<open>e>0\<close> by auto
-      ultimately have "content (cbox a b) < e / 3 * (of_nat (N1 + N2) + 1)"
-        by linarith
-      then have le_e3: "inverse (real (N1 + N2) + 1) * content (cbox a b) \<le> e / 3"
-        unfolding inverse_eq_divide
-        by (auto simp add: field_simps)
-      have ne3: "norm (i (N1 + N2) - s) < e / 3"
-        using N1 by auto
-      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e"
-        apply (rule lem[OF order_trans [OF _ le_e3] ne3 norm_less])
-        apply (rule rsum_diff_bound[OF p[THEN conjunct1]])
-        apply (blast intro: g)
-        done }
-    then show "\<exists>d. gauge d \<and>
-             (\<forall>p. p tagged_division_of cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - s) < e)"
-      by (blast intro: g')
-  qed
-qed
-
-lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified]
-
-
-subsection \<open>Negligible sets.\<close>
-
-definition "negligible (s:: 'a::euclidean_space set) \<longleftrightarrow>
-  (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) (cbox a b))"
-
-
-subsection \<open>Negligibility of hyperplane.\<close>
-
-lemma interval_doublesplit:
-  fixes a :: "'a::euclidean_space"
-  assumes "k \<in> Basis"
-  shows "cbox a b \<inter> {x . \<bar>x\<bullet>k - c\<bar> \<le> (e::real)} =
-    cbox (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) (c - e) else a\<bullet>i) *\<^sub>R i)
-     (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) (c + e) else b\<bullet>i) *\<^sub>R i)"
-proof -
-  have *: "\<And>x c e::real. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
-    by auto
-  have **: "\<And>s P Q. s \<inter> {x. P x \<and> Q x} = (s \<inter> {x. Q x}) \<inter> {x. P x}"
-    by blast
-  show ?thesis
-    unfolding * ** interval_split[OF assms] by (rule refl)
-qed
-
-lemma division_doublesplit:
-  fixes a :: "'a::euclidean_space"
-  assumes "p division_of (cbox a b)"
-    and k: "k \<in> Basis"
-  shows "(\<lambda>l. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e}) ` {l\<in>p. l \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e} \<noteq> {}}
-         division_of  (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> e})"
-proof -
-  have *: "\<And>x c. \<bar>x - c\<bar> \<le> e \<longleftrightarrow> x \<ge> c - e \<and> x \<le> c + e"
-    by auto
-  have **: "\<And>p q p' q'. p division_of q \<Longrightarrow> p = p' \<Longrightarrow> q = q' \<Longrightarrow> p' division_of q'"
-    by auto
-  note division_split(1)[OF assms, where c="c+e",unfolded interval_split[OF k]]
-  note division_split(2)[OF this, where c="c-e" and k=k,OF k]
-  then show ?thesis
-    apply (rule **)
-    subgoal
-      apply (simp add: abs_diff_le_iff field_simps Collect_conj_eq setcompr_eq_image[symmetric])
-      apply (rule equalityI)
-      apply blast
-      apply clarsimp
-      apply (rule_tac x="l \<inter> {x. c + e \<ge> x \<bullet> k}" in exI)
-      apply auto
-      done
-    by (simp add: interval_split k interval_doublesplit)
-qed
-
-lemma content_doublesplit:
-  fixes a :: "'a::euclidean_space"
-  assumes "0 < e"
-    and k: "k \<in> Basis"
-  obtains d where "0 < d" and "content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) < e"
-proof (cases "content (cbox a b) = 0")
-  case True
-  then have ce: "content (cbox a b) < e"
-    by (metis \<open>0 < e\<close>)
-  show ?thesis
-    apply (rule that[of 1])
-    apply simp
-    unfolding interval_doublesplit[OF k]
-    apply (rule le_less_trans[OF content_subset ce])
-    apply (auto simp: interval_doublesplit[symmetric] k)
-    done
-next
-  case False
-  define d where "d = e / 3 / setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
-  note False[unfolded content_eq_0 not_ex not_le, rule_format]
-  then have "\<And>x. x \<in> Basis \<Longrightarrow> b\<bullet>x > a\<bullet>x"
-    by (auto simp add:not_le)
-  then have prod0: "0 < setprod (\<lambda>i. b\<bullet>i - a\<bullet>i) (Basis - {k})"
-    by (force simp add: setprod_pos field_simps)
-  then have "d > 0"
-    using assms
-    by (auto simp add: d_def field_simps)
-  then show ?thesis
-  proof (rule that[of d])
-    have *: "Basis = insert k (Basis - {k})"
-      using k by auto
-    have less_e: "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e"
-    proof -
-      have "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) \<le> 2 * d"
-        by auto
-      also have "\<dots> < e / (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i)"
-        unfolding d_def
-        using assms prod0
-        by (auto simp add: field_simps)
-      finally show "(min (b \<bullet> k) (c + d) - max (a \<bullet> k) (c - d)) * (\<Prod>i\<in>Basis - {k}. b \<bullet> i - a \<bullet> i) < e"
-        unfolding pos_less_divide_eq[OF prod0] .
-    qed
-    show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
-    proof (cases "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = {}")
-      case True
-      then show ?thesis
-        using assms by simp
-    next
-      case False
-      then have
-          "(\<Prod>i\<in>Basis - {k}. interval_upperbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i -
-                interval_lowerbound (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<bullet> i)
-           = (\<Prod>i\<in>Basis - {k}. b\<bullet>i - a\<bullet>i)"
-        by (simp add: box_eq_empty interval_doublesplit[OF k])
-      then show "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
-        unfolding content_def
-        using assms False
-        apply (subst *)
-        apply (subst setprod.insert)
-        apply (simp_all add: interval_doublesplit[OF k] box_eq_empty not_less less_e)
-        done
-    qed
-  qed
-qed
-
-lemma negligible_standard_hyperplane[intro]:
-  fixes k :: "'a::euclidean_space"
-  assumes k: "k \<in> Basis"
-  shows "negligible {x. x\<bullet>k = c}"
-  unfolding negligible_def has_integral
-proof (clarify, goal_cases)
-  case (1 a b e)
-  from this and k obtain d where d: "0 < d" "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
-    by (rule content_doublesplit)
-  let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real"
-  show ?case
-    apply (rule_tac x="\<lambda>x. ball x d" in exI)
-    apply rule
-    apply (rule gauge_ball)
-    apply (rule d)
-  proof (rule, rule)
-    fix p
-    assume p: "p tagged_division_of (cbox a b) \<and> (\<lambda>x. ball x d) fine p"
-    have *: "(\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) =
-      (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)"
-      apply (rule setsum.cong)
-      apply (rule refl)
-      unfolding split_paired_all real_scaleR_def mult_cancel_right split_conv
-      apply cases
-      apply (rule disjI1)
-      apply assumption
-      apply (rule disjI2)
-    proof -
-      fix x l
-      assume as: "(x, l) \<in> p" "?i x \<noteq> 0"
-      then have xk: "x\<bullet>k = c"
-        unfolding indicator_def
-        apply -
-        apply (rule ccontr)
-        apply auto
-        done
-      show "content l = content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
-        apply (rule arg_cong[where f=content])
-        apply (rule set_eqI)
-        apply rule
-        apply rule
-        unfolding mem_Collect_eq
-      proof -
-        fix y
-        assume y: "y \<in> l"
-        note p[THEN conjunct2,unfolded fine_def,rule_format,OF as(1),unfolded split_conv]
-        note this[unfolded subset_eq mem_ball dist_norm,rule_format,OF y]
-        note le_less_trans[OF Basis_le_norm[OF k] this]
-        then show "\<bar>y \<bullet> k - c\<bar> \<le> d"
-          unfolding inner_simps xk by auto
-      qed auto
-    qed
-    note p'= tagged_division_ofD[OF p[THEN conjunct1]] and p''=division_of_tagged_division[OF p[THEN conjunct1]]
-    show "norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R ?i x) - 0) < e"
-      unfolding diff_0_right *
-      unfolding real_scaleR_def real_norm_def
-      apply (subst abs_of_nonneg)
-      apply (rule setsum_nonneg)
-      apply rule
-      unfolding split_paired_all split_conv
-      apply (rule mult_nonneg_nonneg)
-      apply (drule p'(4))
-      apply (erule exE)+
-      apply(rule_tac b=b in back_subst)
-      prefer 2
-      apply (subst(asm) eq_commute)
-      apply assumption
-      apply (subst interval_doublesplit[OF k])
-      apply (rule content_pos_le)
-      apply (rule indicator_pos_le)
-    proof -
-      have "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le>
-        (\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
-        apply (rule setsum_mono)
-        unfolding split_paired_all split_conv
-        apply (rule mult_right_le_one_le)
-        apply (drule p'(4))
-        apply (auto simp add:interval_doublesplit[OF k])
-        done
-      also have "\<dots> < e"
-      proof (subst setsum.over_tagged_division_lemma[OF p[THEN conjunct1]], goal_cases)
-        case prems: (1 u v)
-        have "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content (cbox u v)"
-          unfolding interval_doublesplit[OF k]
-          apply (rule content_subset)
-          unfolding interval_doublesplit[symmetric,OF k]
-          apply auto
-          done
-        then show ?case
-          unfolding prems interval_doublesplit[OF k]
-          by (blast intro: antisym)
-      next
-        have "(\<Sum>l\<in>snd ` p. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
-          setsum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
-        proof (subst (2) setsum.reindex_nontrivial)
-          fix x y assume "x \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}"
-            "x \<noteq> y" and eq: "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
-          then obtain x' y' where "(x', x) \<in> p" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> p" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}"
-            by (auto)
-          from p'(5)[OF \<open>(x', x) \<in> p\<close> \<open>(y', y) \<in> p\<close>] \<open>x \<noteq> y\<close> have "interior (x \<inter> y) = {}"
-            by auto
-          moreover have "interior ((x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<inter> (y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<subseteq> interior (x \<inter> y)"
-            by (auto intro: interior_mono)
-          ultimately have "interior (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = {}"
-            by (auto simp: eq)
-          then show "content (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
-            using p'(4)[OF \<open>(x', x) \<in> p\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int)
-        qed (insert p'(1), auto intro!: setsum.mono_neutral_right)
-        also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` p. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}. content l *\<^sub>R 1::real)"
-          by simp
-        also have "\<dots> \<le> 1 * content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
-          using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]]
-          unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto
-        also have "\<dots> < e"
-          using d(2) by simp
-        finally show "(\<Sum>ka\<in>snd ` p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e" .
-      qed
-      finally show "(\<Sum>(x, ka)\<in>p. content (ka \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" .
-    qed
-  qed
-qed
-
-
-subsection \<open>A technical lemma about "refinement" of division.\<close>
-
-lemma tagged_division_finer:
-  fixes p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
-  assumes "p tagged_division_of (cbox a b)"
-    and "gauge d"
-  obtains q where "q tagged_division_of (cbox a b)"
-    and "d fine q"
-    and "\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q"
-proof -
-  let ?P = "\<lambda>p. p tagged_partial_division_of (cbox a b) \<longrightarrow> gauge d \<longrightarrow>
-    (\<exists>q. q tagged_division_of (\<Union>{k. \<exists>x. (x,k) \<in> p}) \<and> d fine q \<and>
-      (\<forall>(x,k) \<in> p. k \<subseteq> d(x) \<longrightarrow> (x,k) \<in> q))"
-  {
-    have *: "finite p" "p tagged_partial_division_of (cbox a b)"
-      using assms(1)
-      unfolding tagged_division_of_def
-      by auto
-    presume "\<And>p. finite p \<Longrightarrow> ?P p"
-    from this[rule_format,OF * assms(2)] guess q .. note q=this
-    then show ?thesis
-      apply -
-      apply (rule that[of q])
-      unfolding tagged_division_ofD[OF assms(1)]
-      apply auto
-      done
-  }
-  fix p :: "('a::euclidean_space \<times> ('a::euclidean_space set)) set"
-  assume as: "finite p"
-  show "?P p"
-    apply rule
-    apply rule
-    using as
-  proof (induct p)
-    case empty
-    show ?case
-      apply (rule_tac x="{}" in exI)
-      unfolding fine_def
-      apply auto
-      done
-  next
-    case (insert xk p)
-    guess x k using surj_pair[of xk] by (elim exE) note xk=this
-    note tagged_partial_division_subset[OF insert(4) subset_insertI]
-    from insert(3)[OF this insert(5)] guess q1 .. note q1 = conjunctD3[OF this]
-    have *: "\<Union>{l. \<exists>y. (y,l) \<in> insert xk p} = k \<union> \<Union>{l. \<exists>y. (y,l) \<in> p}"
-      unfolding xk by auto
-    note p = tagged_partial_division_ofD[OF insert(4)]
-    from p(4)[unfolded xk, OF insertI1] guess u v by (elim exE) note uv=this
-
-    have "finite {k. \<exists>x. (x, k) \<in> p}"
-      apply (rule finite_subset[of _ "snd ` p"])
-      using p
-      apply safe
-      apply (metis image_iff snd_conv)
-      apply auto
-      done
-    then have int: "interior (cbox u v) \<inter> interior (\<Union>{k. \<exists>x. (x, k) \<in> p}) = {}"
-      apply (rule inter_interior_unions_intervals)
-      apply (rule open_interior)
-      apply (rule_tac[!] ballI)
-      unfolding mem_Collect_eq
-      apply (erule_tac[!] exE)
-      apply (drule p(4)[OF insertI2])
-      apply assumption
-      apply (rule p(5))
-      unfolding uv xk
-      apply (rule insertI1)
-      apply (rule insertI2)
-      apply assumption
-      using insert(2)
-      unfolding uv xk
-      apply auto
-      done
-    show ?case
-    proof (cases "cbox u v \<subseteq> d x")
-      case True
-      then show ?thesis
-        apply (rule_tac x="{(x,cbox u v)} \<union> q1" in exI)
-        apply rule
-        unfolding * uv
-        apply (rule tagged_division_union)
-        apply (rule tagged_division_of_self)
-        apply (rule p[unfolded xk uv] insertI1)+
-        apply (rule q1)
-        apply (rule int)
-        apply rule
-        apply (rule fine_union)
-        apply (subst fine_def)
-        defer
-        apply (rule q1)
-        unfolding Ball_def split_paired_All split_conv
-        apply rule
-        apply rule
-        apply rule
-        apply rule
-        apply (erule insertE)
-        apply (simp add: uv xk)
-        apply (rule UnI2)
-        apply (drule q1(3)[rule_format])
-        unfolding xk uv
-        apply auto
-        done
-    next
-      case False
-      from fine_division_exists[OF assms(2), of u v] guess q2 . note q2=this
-      show ?thesis
-        apply (rule_tac x="q2 \<union> q1" in exI)
-        apply rule
-        unfolding * uv
-        apply (rule tagged_division_union q2 q1 int fine_union)+
-        unfolding Ball_def split_paired_All split_conv
-        apply rule
-        apply (rule fine_union)
-        apply (rule q1 q2)+
-        apply rule
-        apply rule
-        apply rule
-        apply rule
-        apply (erule insertE)
-        apply (rule UnI2)
-        apply (simp add: False uv xk)
-        apply (drule q1(3)[rule_format])
-        using False
-        unfolding xk uv
-        apply auto
-        done
-    qed
-  qed
-qed
-
-
-subsection \<open>Hence the main theorem about negligible sets.\<close>
-
-lemma finite_product_dependent:
-  assumes "finite s"
-    and "\<And>x. x \<in> s \<Longrightarrow> finite (t x)"
-  shows "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
-  using assms
-proof induct
-  case (insert x s)
-  have *: "{(i, j) |i j. i \<in> insert x s \<and> j \<in> t i} =
-    (\<lambda>y. (x,y)) ` (t x) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
-  show ?case
-    unfolding *
-    apply (rule finite_UnI)
-    using insert
-    apply auto
-    done
-qed auto
-
-lemma sum_sum_product:
-  assumes "finite s"
-    and "\<forall>i\<in>s. finite (t i)"
-  shows "setsum (\<lambda>i. setsum (x i) (t i)::real) s =
-    setsum (\<lambda>(i,j). x i j) {(i,j) | i j. i \<in> s \<and> j \<in> t i}"
-  using assms
-proof induct
-  case (insert a s)
-  have *: "{(i, j) |i j. i \<in> insert a s \<and> j \<in> t i} =
-    (\<lambda>y. (a,y)) ` (t a) \<union> {(i, j) |i j. i \<in> s \<and> j \<in> t i}" by auto
-  show ?case
-    unfolding *
-    apply (subst setsum.union_disjoint)
-    unfolding setsum.insert[OF insert(1-2)]
-    prefer 4
-    apply (subst insert(3))
-    unfolding add_right_cancel
-  proof -
-    show "setsum (x a) (t a) = (\<Sum>(xa, y)\<in> Pair a ` t a. x xa y)"
-      apply (subst setsum.reindex)
-      unfolding inj_on_def
-      apply auto
-      done
-    show "finite {(i, j) |i j. i \<in> s \<and> j \<in> t i}"
-      apply (rule finite_product_dependent)
-      using insert
-      apply auto
-      done
-  qed (insert insert, auto)
-qed auto
-
-lemma has_integral_negligible:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "negligible s"
-    and "\<forall>x\<in>(t - s). f x = 0"
-  shows "(f has_integral 0) t"
-proof -
-  presume P: "\<And>f::'b::euclidean_space \<Rightarrow> 'a.
-    \<And>a b. \<forall>x. x \<notin> s \<longrightarrow> f x = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)"
-  let ?f = "(\<lambda>x. if x \<in> t then f x else 0)"
-  show ?thesis
-    apply (rule_tac f="?f" in has_integral_eq)
-    unfolding if_P
-    apply (rule refl)
-    apply (subst has_integral_alt)
-    apply cases
-    apply (subst if_P, assumption)
-    unfolding if_not_P
-  proof -
-    assume "\<exists>a b. t = cbox a b"
-    then guess a b apply - by (erule exE)+ note t = this
-    show "(?f has_integral 0) t"
-      unfolding t
-      apply (rule P)
-      using assms(2)
-      unfolding t
-      apply auto
-      done
-  next
-    show "\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> t then ?f x else 0) has_integral z) (cbox a b) \<and> norm (z - 0) < e)"
-      apply safe
-      apply (rule_tac x=1 in exI)
-      apply rule
-      apply (rule zero_less_one)
-      apply safe
-      apply (rule_tac x=0 in exI)
-      apply rule
-      apply (rule P)
-      using assms(2)
-      apply auto
-      done
-  qed
-next
-  fix f :: "'b \<Rightarrow> 'a"
-  fix a b :: 'b
-  assume assm: "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
-  show "(f has_integral 0) (cbox a b)"
-    unfolding has_integral
-  proof (safe, goal_cases)
-    case prems: (1 e)
-    then have "\<And>n. e / 2 / ((real n+1) * (2 ^ n)) > 0"
-      apply -
-      apply (rule divide_pos_pos)
-      defer
-      apply (rule mult_pos_pos)
-      apply (auto simp add:field_simps)
-      done
-    note assms(1)[unfolded negligible_def has_integral,rule_format,OF this,of a b]
-    note allI[OF this,of "\<lambda>x. x"]
-    from choice[OF this] guess d .. note d=conjunctD2[OF this[rule_format]]
-    show ?case
-      apply (rule_tac x="\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x" in exI)
-    proof safe
-      show "gauge (\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x)"
-        using d(1) unfolding gauge_def by auto
-      fix p
-      assume as: "p tagged_division_of (cbox a b)" "(\<lambda>x. d (nat \<lfloor>norm (f x)\<rfloor>) x) fine p"
-      let ?goal = "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) < e"
-      {
-        presume "p \<noteq> {} \<Longrightarrow> ?goal"
-        then show ?goal
-          apply (cases "p = {}")
-          using prems
-          apply auto
-          done
-      }
-      assume as': "p \<noteq> {}"
-      from real_arch_simple[of "Max((\<lambda>(x,k). norm(f x)) ` p)"] guess N ..
-      then have N: "\<forall>x\<in>(\<lambda>(x, k). norm (f x)) ` p. x \<le> real N"
-        by (meson Max_ge as(1) dual_order.trans finite_imageI tagged_division_of_finite)
-      have "\<forall>i. \<exists>q. q tagged_division_of (cbox a b) \<and> (d i) fine q \<and> (\<forall>(x, k)\<in>p. k \<subseteq> (d i) x \<longrightarrow> (x, k) \<in> q)"
-        by (auto intro: tagged_division_finer[OF as(1) d(1)])
-      from choice[OF this] guess q .. note q=conjunctD3[OF this[rule_format]]
-      have *: "\<And>i. (\<Sum>(x, k)\<in>q i. content k *\<^sub>R indicator s x) \<ge> (0::real)"
-        apply (rule setsum_nonneg)
-        apply safe
-        unfolding real_scaleR_def
-        apply (drule tagged_division_ofD(4)[OF q(1)])
-        apply (auto intro: mult_nonneg_nonneg)
-        done
-      have **: "finite s \<Longrightarrow> finite t \<Longrightarrow> (\<forall>(x,y) \<in> t. (0::real) \<le> g(x,y)) \<Longrightarrow>
-        (\<forall>y\<in>s. \<exists>x. (x,y) \<in> t \<and> f(y) \<le> g(x,y)) \<Longrightarrow> setsum f s \<le> setsum g t" for f g s t
-        apply (rule setsum_le_included[of s t g snd f])
-        prefer 4
-        apply safe
-        apply (erule_tac x=x in ballE)
-        apply (erule exE)
-        apply (rule_tac x="(xa,x)" in bexI)
-        apply auto
-        done
-      have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - 0) \<le> setsum (\<lambda>i. (real i + 1) *
-        norm (setsum (\<lambda>(x,k). content k *\<^sub>R indicator s x :: real) (q i))) {..N+1}"
-        unfolding real_norm_def setsum_right_distrib abs_of_nonneg[OF *] diff_0_right
-        apply (rule order_trans)
-        apply (rule norm_setsum)
-        apply (subst sum_sum_product)
-        prefer 3
-      proof (rule **, safe)
-        show "finite {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i}"
-          apply (rule finite_product_dependent)
-          using q
-          apply auto
-          done
-        fix i a b
-        assume as'': "(a, b) \<in> q i"
-        show "0 \<le> (real i + 1) * (content b *\<^sub>R indicator s a)"
-          unfolding real_scaleR_def
-          using tagged_division_ofD(4)[OF q(1) as'']
-          by (auto intro!: mult_nonneg_nonneg)
-      next
-        fix i :: nat
-        show "finite (q i)"
-          using q by auto
-      next
-        fix x k
-        assume xk: "(x, k) \<in> p"
-        define n where "n = nat \<lfloor>norm (f x)\<rfloor>"
-        have *: "norm (f x) \<in> (\<lambda>(x, k). norm (f x)) ` p"
-          using xk by auto
-        have nfx: "real n \<le> norm (f x)" "norm (f x) \<le> real n + 1"
-          unfolding n_def by auto
-        then have "n \<in> {0..N + 1}"
-          using N[rule_format,OF *] by auto
-        moreover
-        note as(2)[unfolded fine_def,rule_format,OF xk,unfolded split_conv]
-        note q(3)[rule_format,OF xk,unfolded split_conv,rule_format,OF this]
-        note this[unfolded n_def[symmetric]]
-        moreover
-        have "norm (content k *\<^sub>R f x) \<le> (real n + 1) * (content k * indicator s x)"
-        proof (cases "x \<in> s")
-          case False
-          then show ?thesis
-            using assm by auto
-        next
-          case True
-          have *: "content k \<ge> 0"
-            using tagged_division_ofD(4)[OF as(1) xk] by auto
-          moreover
-          have "content k * norm (f x) \<le> content k * (real n + 1)"
-            apply (rule mult_mono)
-            using nfx *
-            apply auto
-            done
-          ultimately
-          show ?thesis
-            unfolding abs_mult
-            using nfx True
-            by (auto simp add: field_simps)
-        qed
-        ultimately show "\<exists>y. (y, x, k) \<in> {(i, j) |i j. i \<in> {..N + 1} \<and> j \<in> q i} \<and> norm (content k *\<^sub>R f x) \<le>
-          (real y + 1) * (content k *\<^sub>R indicator s x)"
-          apply (rule_tac x=n in exI)
-          apply safe
-          apply (rule_tac x=n in exI)
-          apply (rule_tac x="(x,k)" in exI)
-          apply safe
-          apply auto
-          done
-      qed (insert as, auto)
-      also have "\<dots> \<le> setsum (\<lambda>i. e / 2 / 2 ^ i) {..N+1}"
-      proof (rule setsum_mono, goal_cases)
-        case (1 i)
-        then show ?case
-          apply (subst mult.commute, subst pos_le_divide_eq[symmetric])
-          using d(2)[rule_format, of "q i" i]
-          using q[rule_format]
-          apply (auto simp add: field_simps)
-          done
-      qed
-      also have "\<dots> < e * inverse 2 * 2"
-        unfolding divide_inverse setsum_right_distrib[symmetric]
-        apply (rule mult_strict_left_mono)
-        unfolding power_inverse [symmetric] lessThan_Suc_atMost[symmetric]
-        apply (subst geometric_sum)
-        using prems
-        apply auto
-        done
-      finally show "?goal" by auto
-    qed
-  qed
-qed
-
-lemma has_integral_spike:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector"
-  assumes "negligible s"
-    and "(\<forall>x\<in>(t - s). g x = f x)"
-    and "(f has_integral y) t"
-  shows "(g has_integral y) t"
-proof -
-  {
-    fix a b :: 'b
-    fix f g :: "'b \<Rightarrow> 'a"
-    fix y :: 'a
-    assume as: "\<forall>x \<in> cbox a b - s. g x = f x" "(f has_integral y) (cbox a b)"
-    have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)"
-      apply (rule has_integral_add[OF as(2)])
-      apply (rule has_integral_negligible[OF assms(1)])
-      using as
-      apply auto
-      done
-    then have "(g has_integral y) (cbox a b)"
-      by auto
-  } note * = this
-  show ?thesis
-    apply (subst has_integral_alt)
-    using assms(2-)
-    apply -
-    apply (rule cond_cases)
-    apply safe
-    apply (rule *)
-    apply assumption+
-    apply (subst(asm) has_integral_alt)
-    unfolding if_not_P
-    apply (erule_tac x=e in allE)
-    apply safe
-    apply (rule_tac x=B in exI)
-    apply safe
-    apply (erule_tac x=a in allE)
-    apply (erule_tac x=b in allE)
-    apply safe
-    apply (rule_tac x=z in exI)
-    apply safe
-    apply (rule *[where fa2="\<lambda>x. if x\<in>t then f x else 0"])
-    apply auto
-    done
-qed
-
-lemma has_integral_spike_eq:
-  assumes "negligible s"
-    and "\<forall>x\<in>(t - s). g x = f x"
-  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
-  apply rule
-  apply (rule_tac[!] has_integral_spike[OF assms(1)])
-  using assms(2)
-  apply auto
-  done
-
-lemma integrable_spike:
-  assumes "negligible s"
-    and "\<forall>x\<in>(t - s). g x = f x"
-    and "f integrable_on t"
-  shows "g integrable_on  t"
-  using assms
-  unfolding integrable_on_def
-  apply -
-  apply (erule exE)
-  apply rule
-  apply (rule has_integral_spike)
-  apply fastforce+
-  done
-
-lemma integral_spike:
-  assumes "negligible s"
-    and "\<forall>x\<in>(t - s). g x = f x"
-  shows "integral t f = integral t g"
-  using has_integral_spike_eq[OF assms] by (simp add: integral_def integrable_on_def)
-
-
-subsection \<open>Some other trivialities about negligible sets.\<close>
-
-lemma negligible_subset[intro]:
-  assumes "negligible s"
-    and "t \<subseteq> s"
-  shows "negligible t"
-  unfolding negligible_def
-proof (safe, goal_cases)
-  case (1 a b)
-  show ?case
-    using assms(1)[unfolded negligible_def,rule_format,of a b]
-    apply -
-    apply (rule has_integral_spike[OF assms(1)])
-    defer
-    apply assumption
-    using assms(2)
-    unfolding indicator_def
-    apply auto
-    done
-qed
-
-lemma negligible_diff[intro?]:
-  assumes "negligible s"
-  shows "negligible (s - t)"
-  using assms by auto
-
-lemma negligible_Int:
-  assumes "negligible s \<or> negligible t"
-  shows "negligible (s \<inter> t)"
-  using assms by auto
-
-lemma negligible_Un:
-  assumes "negligible s"
-    and "negligible t"
-  shows "negligible (s \<union> t)"
-  unfolding negligible_def
-proof (safe, goal_cases)
-  case (1 a b)
-  note assm = assms[unfolded negligible_def,rule_format,of a b]
-  then show ?case
-    apply (subst has_integral_spike_eq[OF assms(2)])
-    defer
-    apply assumption
-    unfolding indicator_def
-    apply auto
-    done
-qed
-
-lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t"
-  using negligible_Un by auto
-
-lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}"
-  using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] by auto
-
-lemma negligible_insert[simp]: "negligible (insert a s) \<longleftrightarrow> negligible s"
-  apply (subst insert_is_Un)
-  unfolding negligible_Un_eq
-  apply auto
-  done
-
-lemma negligible_empty[iff]: "negligible {}"
-  by auto
-
-lemma negligible_finite[intro]:
-  assumes "finite s"
-  shows "negligible s"
-  using assms by (induct s) auto
-
-lemma negligible_Union[intro]:
-  assumes "finite s"
-    and "\<forall>t\<in>s. negligible t"
-  shows "negligible(\<Union>s)"
-  using assms by induct auto
-
-lemma negligible:
-  "negligible s \<longleftrightarrow> (\<forall>t::('a::euclidean_space) set. ((indicator s::'a\<Rightarrow>real) has_integral 0) t)"
-  apply safe
-  defer
-  apply (subst negligible_def)
-proof -
-  fix t :: "'a set"
-  assume as: "negligible s"
-  have *: "(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)"
-    by auto
-  show "((indicator s::'a\<Rightarrow>real) has_integral 0) t"
-    apply (subst has_integral_alt)
-    apply cases
-    apply (subst if_P,assumption)
-    unfolding if_not_P
-    apply safe
-    apply (rule as[unfolded negligible_def,rule_format])
-    apply (rule_tac x=1 in exI)
-    apply safe
-    apply (rule zero_less_one)
-    apply (rule_tac x=0 in exI)
-    using negligible_subset[OF as,of "s \<inter> t"]
-    unfolding negligible_def indicator_def [abs_def]
-    unfolding *
-    apply auto
-    done
-qed auto
-
-
-subsection \<open>Finite case of the spike theorem is quite commonly needed.\<close>
-
-lemma has_integral_spike_finite:
-  assumes "finite s"
-    and "\<forall>x\<in>t-s. g x = f x"
-    and "(f has_integral y) t"
-  shows "(g has_integral y) t"
-  apply (rule has_integral_spike)
-  using assms
-  apply auto
-  done
-
-lemma has_integral_spike_finite_eq:
-  assumes "finite s"
-    and "\<forall>x\<in>t-s. g x = f x"
-  shows "((f has_integral y) t \<longleftrightarrow> (g has_integral y) t)"
-  apply rule
-  apply (rule_tac[!] has_integral_spike_finite)
-  using assms
-  apply auto
-  done
-
-lemma integrable_spike_finite:
-  assumes "finite s"
-    and "\<forall>x\<in>t-s. g x = f x"
-    and "f integrable_on t"
-  shows "g integrable_on  t"
-  using assms
-  unfolding integrable_on_def
-  apply safe
-  apply (rule_tac x=y in exI)
-  apply (rule has_integral_spike_finite)
-  apply auto
-  done
-
-
-subsection \<open>In particular, the boundary of an interval is negligible.\<close>
-
-lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)"
-proof -
-  let ?A = "\<Union>((\<lambda>k. {x. x\<bullet>k = a\<bullet>k} \<union> {x::'a. x\<bullet>k = b\<bullet>k}) ` Basis)"
-  have "cbox a b - box a b \<subseteq> ?A"
-    apply rule unfolding Diff_iff mem_box
-    apply simp
-    apply(erule conjE bexE)+
-    apply(rule_tac x=i in bexI)
-    apply auto
-    done
-  then show ?thesis
-    apply -
-    apply (rule negligible_subset[of ?A])
-    apply (rule negligible_Union[OF finite_imageI])
-    apply auto
-    done
-qed
-
-lemma has_integral_spike_interior:
-  assumes "\<forall>x\<in>box a b. g x = f x"
-    and "(f has_integral y) (cbox a b)"
-  shows "(g has_integral y) (cbox a b)"
-  apply (rule has_integral_spike[OF negligible_frontier_interval _ assms(2)])
-  using assms(1)
-  apply auto
-  done
-
-lemma has_integral_spike_interior_eq:
-  assumes "\<forall>x\<in>box a b. g x = f x"
-  shows "(f has_integral y) (cbox a b) \<longleftrightarrow> (g has_integral y) (cbox a b)"
-  apply rule
-  apply (rule_tac[!] has_integral_spike_interior)
-  using assms
-  apply auto
-  done
-
-lemma integrable_spike_interior:
-  assumes "\<forall>x\<in>box a b. g x = f x"
-    and "f integrable_on cbox a b"
-  shows "g integrable_on cbox a b"
-  using assms
-  unfolding integrable_on_def
-  using has_integral_spike_interior[OF assms(1)]
-  by auto
-
-
-subsection \<open>Integrability of continuous functions.\<close>
-
-lemma operative_approximable:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "0 \<le> e"
-  shows "comm_monoid.operative op \<and> True (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)"
-  unfolding comm_monoid.operative_def[OF comm_monoid_and]
-proof safe
-  fix a b :: 'b
-  show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
-    if "content (cbox a b) = 0"
-    apply (rule_tac x=f in exI)
-    using assms that
-    apply (auto intro!: integrable_on_null)
-    done
-  {
-    fix c g
-    fix k :: 'b
-    assume as: "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b"
-    assume k: "k \<in> Basis"
-    show "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
-      "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
-      apply (rule_tac[!] x=g in exI)
-      using as(1) integrable_split[OF as(2) k]
-      apply auto
-      done
-  }
-  fix c k g1 g2
-  assume as: "\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e" "g1 integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
-    "\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e" "g2 integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
-  assume k: "k \<in> Basis"
-  let ?g = "\<lambda>x. if x\<bullet>k = c then f x else if x\<bullet>k \<le> c then g1 x else g2 x"
-  show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
-    apply (rule_tac x="?g" in exI)
-    apply safe
-  proof goal_cases
-    case (1 x)
-    then show ?case
-      apply -
-      apply (cases "x\<bullet>k=c")
-      apply (case_tac "x\<bullet>k < c")
-      using as assms
-      apply auto
-      done
-  next
-    case 2
-    presume "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
-      and "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
-    then guess h1 h2 unfolding integrable_on_def by auto
-    from has_integral_split[OF this k] show ?case
-      unfolding integrable_on_def by auto
-  next
-    show "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
-      apply(rule_tac[!] integrable_spike[OF negligible_standard_hyperplane[of k c]])
-      using k as(2,4)
-      apply auto
-      done
-  qed
-qed
-
-lemma comm_monoid_set_F_and: "comm_monoid_set.F op \<and> True f s \<longleftrightarrow> (finite s \<longrightarrow> (\<forall>x\<in>s. f x))"
-proof -
-  interpret bool: comm_monoid_set "op \<and>" True
-    proof qed auto
-  show ?thesis
-    by (induction s rule: infinite_finite_induct) auto
-qed
-
-lemma approximable_on_division:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "0 \<le> e"
-    and "d division_of (cbox a b)"
-    and "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
-  obtains g where "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b"
-proof -
-  note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_approximable[OF assms(1)] assms(2)]
-  from assms(3) this[unfolded comm_monoid_set_F_and, of f] division_of_finite[OF assms(2)]
-  guess g by auto
-  then show thesis
-    apply -
-    apply (rule that[of g])
-    apply auto
-    done
-qed
-
-lemma integrable_continuous:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "continuous_on (cbox a b) f"
-  shows "f integrable_on cbox a b"
-proof (rule integrable_uniform_limit, safe)
-  fix e :: real
-  assume e: "e > 0"
-  from compact_uniformly_continuous[OF assms compact_cbox,unfolded uniformly_continuous_on_def,rule_format,OF e] guess d ..
-  note d=conjunctD2[OF this,rule_format]
-  from fine_division_exists[OF gauge_ball[OF d(1)], of a b] guess p . note p=this
-  note p' = tagged_division_ofD[OF p(1)]
-  have *: "\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i"
-  proof (safe, unfold snd_conv)
-    fix x l
-    assume as: "(x, l) \<in> p"
-    from p'(4)[OF this] guess a b by (elim exE) note l=this
-    show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l"
-      apply (rule_tac x="\<lambda>y. f x" in exI)
-    proof safe
-      show "(\<lambda>y. f x) integrable_on l"
-        unfolding integrable_on_def l
-        apply rule
-        apply (rule has_integral_const)
-        done
-      fix y
-      assume y: "y \<in> l"
-      note fineD[OF p(2) as,unfolded subset_eq,rule_format,OF this]
-      note d(2)[OF _ _ this[unfolded mem_ball]]
-      then show "norm (f y - f x) \<le> e"
-        using y p'(2-3)[OF as] unfolding dist_norm l norm_minus_commute by fastforce
-    qed
-  qed
-  from e have "e \<ge> 0"
-    by auto
-  from approximable_on_division[OF this division_of_tagged_division[OF p(1)] *] guess g .
-  then show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b"
-    by auto
-qed
-
-lemma integrable_continuous_real:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a .. b} f"
-  shows "f integrable_on {a .. b}"
-  by (metis assms box_real(2) integrable_continuous)
-
-subsection \<open>Specialization of additivity to one dimension.\<close>
-
-subsection \<open>Special case of additivity we need for the FTC.\<close>
-
-lemma additive_tagged_division_1:
-  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b"
-    and "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). f(Sup k) - f(Inf k)) p = f b - f a"
-proof -
-  let ?f = "(\<lambda>k::(real) set. if k = {} then 0 else f(interval_upperbound k) - f(interval_lowerbound k))"
-  have ***: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i"
-    using assms by auto
-  have *: "add.operative ?f"
-    unfolding add.operative_1_lt box_eq_empty
-    by auto
-  have **: "cbox a b \<noteq> {}"
-    using assms(1) by auto
-  note setsum.operative_tagged_division[OF * assms(2)[simplified box_real[symmetric]]]
-  note * = this[unfolded if_not_P[OF **] interval_bounds[OF ***],symmetric]
-  show ?thesis
-    unfolding *
-    apply (rule setsum.cong)
-    unfolding split_paired_all split_conv
-    using assms(2)
-    apply auto
-    done
-qed
-
-
-subsection \<open>A useful lemma allowing us to factor out the content size.\<close>
-
-lemma has_integral_factor_content:
-  "(f has_integral i) (cbox a b) \<longleftrightarrow>
-    (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))"
-proof (cases "content (cbox a b) = 0")
-  case True
-  show ?thesis
-    unfolding has_integral_null_eq[OF True]
-    apply safe
-    apply (rule, rule, rule gauge_trivial, safe)
-    unfolding setsum_content_null[OF True] True
-    defer
-    apply (erule_tac x=1 in allE)
-    apply safe
-    defer
-    apply (rule fine_division_exists[of _ a b])
-    apply assumption
-    apply (erule_tac x=p in allE)
-    unfolding setsum_content_null[OF True]
-    apply auto
-    done
-next
-  case False
-  note F = this[unfolded content_lt_nz[symmetric]]
-  let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and>
-    (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)"
-  show ?thesis
-    apply (subst has_integral)
-  proof safe
-    fix e :: real
-    assume e: "e > 0"
-    {
-      assume "\<forall>e>0. ?P e op <"
-      then show "?P (e * content (cbox a b)) op \<le>"
-        apply (erule_tac x="e * content (cbox a b)" in allE)
-        apply (erule impE)
-        defer
-        apply (erule exE,rule_tac x=d in exI)
-        using F e
-        apply (auto simp add:field_simps)
-        done
-    }
-    {
-      assume "\<forall>e>0. ?P (e * content (cbox a b)) op \<le>"
-      then show "?P e op <"
-        apply (erule_tac x="e / 2 / content (cbox a b)" in allE)
-        apply (erule impE)
-        defer
-        apply (erule exE,rule_tac x=d in exI)
-        using F e
-        apply (auto simp add: field_simps)
-        done
-    }
-  qed
-qed
-
-lemma has_integral_factor_content_real:
-  "(f has_integral i) {a .. b::real} \<longleftrightarrow>
-    (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b}  \<and> d fine p \<longrightarrow>
-      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a .. b} ))"
-  unfolding box_real[symmetric]
-  by (rule has_integral_factor_content)
-
-
-subsection \<open>Fundamental theorem of calculus.\<close>
-
-lemma interval_bounds_real:
-  fixes q b :: real
-  assumes "a \<le> b"
-  shows "Sup {a..b} = b"
-    and "Inf {a..b} = a"
-  using assms by auto
-
-lemma fundamental_theorem_of_calculus:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "a \<le> b"
-    and "\<forall>x\<in>{a .. b}. (f has_vector_derivative f' x) (at x within {a .. b})"
-  shows "(f' has_integral (f b - f a)) {a .. b}"
-  unfolding has_integral_factor_content box_real[symmetric]
-proof safe
-  fix e :: real
-  assume e: "e > 0"
-  note assm = assms(2)[unfolded has_vector_derivative_def has_derivative_within_alt]
-  have *: "\<And>P Q. \<forall>x\<in>{a .. b}. P x \<and> (\<forall>e>0. \<exists>d>0. Q x e d) \<Longrightarrow> \<forall>x. \<exists>(d::real)>0. x\<in>{a .. b} \<longrightarrow> Q x e d"
-    using e by blast
-  note this[OF assm,unfolded gauge_existence_lemma]
-  from choice[OF this,unfolded Ball_def[symmetric]] guess d ..
-  note d=conjunctD2[OF this[rule_format],rule_format]
-  show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b))"
-    apply (rule_tac x="\<lambda>x. ball x (d x)" in exI)
-    apply safe
-    apply (rule gauge_ball_dependent)
-    apply rule
-    apply (rule d(1))
-  proof -
-    fix p
-    assume as: "p tagged_division_of cbox a b" "(\<lambda>x. ball x (d x)) fine p"
-    show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)"
-      unfolding content_real[OF assms(1), simplified box_real[symmetric]] additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of f,symmetric]
-      unfolding additive_tagged_division_1[OF assms(1) as(1)[simplified box_real],of "\<lambda>x. x",symmetric]
-      unfolding setsum_right_distrib
-      defer
-      unfolding setsum_subtractf[symmetric]
-    proof (rule setsum_norm_le,safe)
-      fix x k
-      assume "(x, k) \<in> p"
-      note xk = tagged_division_ofD(2-4)[OF as(1) this]
-      from this(3) guess u v by (elim exE) note k=this
-      have *: "u \<le> v"
-        using xk unfolding k by auto
-      have ball: "\<forall>xa\<in>k. xa \<in> ball x (d x)"
-        using as(2)[unfolded fine_def,rule_format,OF \<open>(x,k)\<in>p\<close>,unfolded split_conv subset_eq] .
-      have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le>
-        norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)"
-        apply (rule order_trans[OF _ norm_triangle_ineq4])
-        apply (rule eq_refl)
-        apply (rule arg_cong[where f=norm])
-        unfolding scaleR_diff_left
-        apply (auto simp add:algebra_simps)
-        done
-      also have "\<dots> \<le> e * norm (u - x) + e * norm (v - x)"
-        apply (rule add_mono)
-        apply (rule d(2)[of "x" "u",unfolded o_def])
-        prefer 4
-        apply (rule d(2)[of "x" "v",unfolded o_def])
-        using ball[rule_format,of u] ball[rule_format,of v]
-        using xk(1-2)
-        unfolding k subset_eq
-        apply (auto simp add:dist_real_def)
-        done
-      also have "\<dots> \<le> e * (Sup k - Inf k)"
-        unfolding k interval_bounds_real[OF *]
-        using xk(1)
-        unfolding k
-        by (auto simp add: dist_real_def field_simps)
-      finally show "norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le>
-        e * (Sup k - Inf k)"
-        unfolding box_real k interval_bounds_real[OF *] content_real[OF *]
-          interval_upperbound_real interval_lowerbound_real
-          .
-    qed
-  qed
-qed
-
-lemma ident_has_integral:
-  fixes a::real
-  assumes "a \<le> b"
-  shows "((\<lambda>x. x) has_integral (b\<^sup>2 - a\<^sup>2) / 2) {a..b}"
-proof -
-  have "((\<lambda>x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}"
-    apply (rule fundamental_theorem_of_calculus [OF assms], clarify)
-    unfolding power2_eq_square
-    by (rule derivative_eq_intros | simp)+
-  then show ?thesis
-    by (simp add: field_simps)
-qed
-
-lemma integral_ident [simp]:
-  fixes a::real
-  assumes "a \<le> b"
-  shows "integral {a..b} (\<lambda>x. x) = (if a \<le> b then (b\<^sup>2 - a\<^sup>2) / 2 else 0)"
-using ident_has_integral integral_unique by fastforce
-
-lemma ident_integrable_on:
-  fixes a::real
-  shows "(\<lambda>x. x) integrable_on {a..b}"
-by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral)
-
-
-subsection \<open>Taylor series expansion\<close>
-
-lemma (in bounded_bilinear) setsum_prod_derivatives_has_vector_derivative:
-  assumes "p>0"
-  and f0: "Df 0 = f"
-  and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
-    (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})"
-  and g0: "Dg 0 = g"
-  and Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
-    (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})"
-  and ivl: "a \<le> t" "t \<le> b"
-  shows "((\<lambda>t. \<Sum>i<p. (-1)^i *\<^sub>R prod (Df i t) (Dg (p - Suc i) t))
-    has_vector_derivative
-      prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t))
-    (at t within {a .. b})"
-  using assms
-proof cases
-  assume p: "p \<noteq> 1"
-  define p' where "p' = p - 2"
-  from assms p have p': "{..<p} = {..Suc p'}" "p = Suc (Suc p')"
-    by (auto simp: p'_def)
-  have *: "\<And>i. i \<le> p' \<Longrightarrow> Suc (Suc p' - i) = (Suc (Suc p') - i)"
-    by auto
-  let ?f = "\<lambda>i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))"
-  have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
-    prod (Df (Suc i) t) (Dg (p - Suc i) t))) =
-    (\<Sum>i\<le>(Suc p'). ?f i - ?f (Suc i))"
-    by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost)
-  also note setsum_telescope
-  finally
-  have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) +
-    prod (Df (Suc i) t) (Dg (p - Suc i) t)))
-    = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)"
-    unfolding p'[symmetric]
-    by (simp add: assms)
-  thus ?thesis
-    using assms
-    by (auto intro!: derivative_eq_intros has_vector_derivative)
-qed (auto intro!: derivative_eq_intros has_vector_derivative)
-
-lemma
-  fixes f::"real\<Rightarrow>'a::banach"
-  assumes "p>0"
-  and f0: "Df 0 = f"
-  and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
-    (Df m has_vector_derivative Df (Suc m) t) (at t within {a .. b})"
-  and ivl: "a \<le> b"
-  defines "i \<equiv> \<lambda>x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x"
-  shows taylor_has_integral:
-    "(i has_integral f b - (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a)) {a..b}"
-  and taylor_integral:
-    "f b = (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a) + integral {a..b} i"
-  and taylor_integrable:
-    "i integrable_on {a .. b}"
-proof goal_cases
-  case 1
-  interpret bounded_bilinear "scaleR::real\<Rightarrow>'a\<Rightarrow>'a"
-    by (rule bounded_bilinear_scaleR)
-  define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s
-  define Dg where [abs_def]:
-    "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s
-  have g0: "Dg 0 = g"
-    using \<open>p > 0\<close>
-    by (auto simp add: Dg_def divide_simps g_def split: if_split_asm)
-  {
-    fix m
-    assume "p > Suc m"
-    hence "p - Suc m = Suc (p - Suc (Suc m))"
-      by auto
-    hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)"
-      by auto
-  } note fact_eq = this
-  have Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
-    (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a .. b})"
-    unfolding Dg_def
-    by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps)
-  let ?sum = "\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t"
-  from setsum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df,
-      OF \<open>p > 0\<close> g0 Dg f0 Df]
-  have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow>
-    (?sum has_vector_derivative
-      g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})"
-    by auto
-  from fundamental_theorem_of_calculus[rule_format, OF \<open>a \<le> b\<close> deriv]
-  have "(i has_integral ?sum b - ?sum a) {a .. b}"
-    by (simp add: i_def g_def Dg_def)
-  also
-  have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)"
-    and "{..<p} \<inter> {i. p = Suc i} = {p - 1}"
-    for p'
-    using \<open>p > 0\<close>
-    by (auto simp: power_mult_distrib[symmetric])
-  then have "?sum b = f b"
-    using Suc_pred'[OF \<open>p > 0\<close>]
-    by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib
-        cond_application_beta setsum.If_cases f0)
-  also
-  have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
-  proof safe
-    fix x
-    assume "x < p"
-    thus "x \<in> (\<lambda>x. p - x - 1) ` {..<p}"
-      by (auto intro!: image_eqI[where x = "p - x - 1"])
-  qed simp
-  from _ this
-  have "?sum a = (\<Sum>i<p. ((b - a) ^ i / fact i) *\<^sub>R Df i a)"
-    by (rule setsum.reindex_cong) (auto simp add: inj_on_def Dg_def one)
-  finally show c: ?case .
-  case 2 show ?case using c integral_unique by force
-  case 3 show ?case using c by force
-qed
-
-
-subsection \<open>Attempt a systematic general set of "offset" results for components.\<close>
-
-lemma gauge_modify:
-  assumes "(\<forall>s. open s \<longrightarrow> open {x. f(x) \<in> s})" "gauge d"
-  shows "gauge (\<lambda>x. {y. f y \<in> d (f x)})"
-  using assms
-  unfolding gauge_def
-  apply safe
-  defer
-  apply (erule_tac x="f x" in allE)
-  apply (erule_tac x="d (f x)" in allE)
-  apply auto
-  done
-
-
-subsection \<open>Only need trivial subintervals if the interval itself is trivial.\<close>
-
-lemma division_of_nontrivial:
-  fixes s :: "'a::euclidean_space set set"
-  assumes "s division_of (cbox a b)"
-    and "content (cbox a b) \<noteq> 0"
-  shows "{k. k \<in> s \<and> content k \<noteq> 0} division_of (cbox a b)"
-  using assms(1)
-  apply -
-proof (induct "card s" arbitrary: s rule: nat_less_induct)
-  fix s::"'a set set"
-  assume assm: "s division_of (cbox a b)"
-    "\<forall>m<card s. \<forall>x. m = card x \<longrightarrow>
-      x division_of (cbox a b) \<longrightarrow> {k \<in> x. content k \<noteq> 0} division_of (cbox a b)"
-  note s = division_ofD[OF assm(1)]
-  let ?thesis = "{k \<in> s. content k \<noteq> 0} division_of (cbox a b)"
-  {
-    presume *: "{k \<in> s. content k \<noteq> 0} \<noteq> s \<Longrightarrow> ?thesis"
-    show ?thesis
-      apply cases
-      defer
-      apply (rule *)
-      apply assumption
-      using assm(1)
-      apply auto
-      done
-  }
-  assume noteq: "{k \<in> s. content k \<noteq> 0} \<noteq> s"
-  then obtain k where k: "k \<in> s" "content k = 0"
-    by auto
-  from s(4)[OF k(1)] guess c d by (elim exE) note k=k this
-  from k have "card s > 0"
-    unfolding card_gt_0_iff using assm(1) by auto
-  then have card: "card (s - {k}) < card s"
-    using assm(1) k(1)
-    apply (subst card_Diff_singleton_if)
-    apply auto
-    done
-  have *: "closed (\<Union>(s - {k}))"
-    apply (rule closed_Union)
-    defer
-    apply rule
-    apply (drule DiffD1,drule s(4))
-    using assm(1)
-    apply auto
-    done
-  have "k \<subseteq> \<Union>(s - {k})"
-    apply safe
-    apply (rule *[unfolded closed_limpt,rule_format])
-    unfolding islimpt_approachable
-  proof safe
-    fix x
-    fix e :: real
-    assume as: "x \<in> k" "e > 0"
-    from k(2)[unfolded k content_eq_0] guess i ..
-    then have i:"c\<bullet>i = d\<bullet>i" "i\<in>Basis"
-      using s(3)[OF k(1),unfolded k] unfolding box_ne_empty by auto
-    then have xi: "x\<bullet>i = d\<bullet>i"
-      using as unfolding k mem_box by (metis antisym)
-    define y where "y = (\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i) / 2 then c\<bullet>i +
-      min e (b\<bullet>i - c\<bullet>i) / 2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i) / 2 else x\<bullet>j) *\<^sub>R j)"
-    show "\<exists>x'\<in>\<Union>(s - {k}). x' \<noteq> x \<and> dist x' x < e"
-      apply (rule_tac x=y in bexI)
-    proof
-      have "d \<in> cbox c d"
-        using s(3)[OF k(1)]
-        unfolding k box_eq_empty mem_box
-        by (fastforce simp add: not_less)
-      then have "d \<in> cbox a b"
-        using s(2)[OF k(1)]
-        unfolding k
-        by auto
-      note di = this[unfolded mem_box,THEN bspec[where x=i]]
-      then have xyi: "y\<bullet>i \<noteq> x\<bullet>i"
-        unfolding y_def i xi
-        using as(2) assms(2)[unfolded content_eq_0] i(2)
-        by (auto elim!: ballE[of _ _ i])
-      then show "y \<noteq> x"
-        unfolding euclidean_eq_iff[where 'a='a] using i by auto
-      have *: "Basis = insert i (Basis - {i})"
-        using i by auto
-      have "norm (y - x) < e + setsum (\<lambda>i. 0) Basis"
-        apply (rule le_less_trans[OF norm_le_l1])
-        apply (subst *)
-        apply (subst setsum.insert)
-        prefer 3
-        apply (rule add_less_le_mono)
-      proof -
-        show "\<bar>(y - x) \<bullet> i\<bar> < e"
-          using di as(2) y_def i xi by (auto simp: inner_simps)
-        show "(\<Sum>i\<in>Basis - {i}. \<bar>(y - x) \<bullet> i\<bar>) \<le> (\<Sum>i\<in>Basis. 0)"
-          unfolding y_def by (auto simp: inner_simps)
-      qed auto
-      then show "dist y x < e"
-        unfolding dist_norm by auto
-      have "y \<notin> k"
-        unfolding k mem_box
-        apply rule
-        apply (erule_tac x=i in ballE)
-        using xyi k i xi
-        apply auto
-        done
-      moreover
-      have "y \<in> \<Union>s"
-        using set_rev_mp[OF as(1) s(2)[OF k(1)]] as(2) di i
-        unfolding s mem_box y_def
-        by (auto simp: field_simps elim!: ballE[of _ _ i])
-      ultimately
-      show "y \<in> \<Union>(s - {k})" by auto
-    qed
-  qed
-  then have "\<Union>(s - {k}) = cbox a b"
-    unfolding s(6)[symmetric] by auto
-  then have  "{ka \<in> s - {k}. content ka \<noteq> 0} division_of (cbox a b)"
-    apply -
-    apply (rule assm(2)[rule_format,OF card refl])
-    apply (rule division_ofI)
-    defer
-    apply (rule_tac[1-4] s)
-    using assm(1)
-    apply auto
-    done
-  moreover
-  have "{ka \<in> s - {k}. content ka \<noteq> 0} = {k \<in> s. content k \<noteq> 0}"
-    using k by auto
-  ultimately show ?thesis by auto
-qed
-
-
-subsection \<open>Integrability on subintervals.\<close>
-
-lemma operative_integrable:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  shows "comm_monoid.operative op \<and> True (\<lambda>i. f integrable_on i)"
-  unfolding comm_monoid.operative_def[OF comm_monoid_and]
-  apply safe
-  apply (subst integrable_on_def)
-  unfolding has_integral_null_eq
-  apply (rule, rule refl)
-  apply (rule, assumption, assumption)+
-  unfolding integrable_on_def
-  by (auto intro!: has_integral_split)
-
-lemma integrable_subinterval:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on cbox a b"
-    and "cbox c d \<subseteq> cbox a b"
-  shows "f integrable_on cbox c d"
-  apply (cases "cbox c d = {}")
-  defer
-  apply (rule partial_division_extend_1[OF assms(2)],assumption)
-  using comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable,symmetric,of _ _ _ f] assms(1)
-  apply (auto simp: comm_monoid_set_F_and)
-  done
-
-lemma integrable_subinterval_real:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "f integrable_on {a .. b}"
-    and "{c .. d} \<subseteq> {a .. b}"
-  shows "f integrable_on {c .. d}"
-  by (metis assms(1) assms(2) box_real(2) integrable_subinterval)
-
-
-subsection \<open>Combining adjacent intervals in 1 dimension.\<close>
-
-lemma has_integral_combine:
-  fixes a b c :: real
-  assumes "a \<le> c"
-    and "c \<le> b"
-    and "(f has_integral i) {a .. c}"
-    and "(f has_integral (j::'a::banach)) {c .. b}"
-  shows "(f has_integral (i + j)) {a .. b}"
-proof -
-  note operative_integral[of f, unfolded comm_monoid.operative_1_le[OF add.comm_monoid_lift_option]]
-  note conjunctD2[OF this,rule_format]
-  note * = this(2)[OF conjI[OF assms(1-2)],unfolded if_P[OF assms(3)]]
-  then have "f integrable_on cbox a b"
-    apply -
-    apply (rule ccontr)
-    apply (subst(asm) if_P)
-    defer
-    apply (subst(asm) if_P)
-    using assms(3-)
-    apply auto
-    done
-  with *
-  show ?thesis
-    apply -
-    apply (subst(asm) if_P)
-    defer
-    apply (subst(asm) if_P)
-    defer
-    apply (subst(asm) if_P)
-    using assms(3-)
-    apply (auto simp add: integrable_on_def integral_unique)
-    done
-qed
-
-lemma integral_combine:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "a \<le> c"
-    and "c \<le> b"
-    and "f integrable_on {a .. b}"
-  shows "integral {a .. c} f + integral {c .. b} f = integral {a .. b} f"
-  apply (rule integral_unique[symmetric])
-  apply (rule has_integral_combine[OF assms(1-2)])
-  apply (metis assms(2) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel2 monoid_add_class.add.left_neutral)
-  by (metis assms(1) assms(3) atLeastatMost_subset_iff box_real(2) content_pos_le content_real_eq_0 integrable_integral integrable_subinterval le_add_same_cancel1 monoid_add_class.add.right_neutral)
-
-lemma integrable_combine:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "a \<le> c"
-    and "c \<le> b"
-    and "f integrable_on {a .. c}"
-    and "f integrable_on {c .. b}"
-  shows "f integrable_on {a .. b}"
-  using assms
-  unfolding integrable_on_def
-  by (fastforce intro!:has_integral_combine)
-
-
-subsection \<open>Reduce integrability to "local" integrability.\<close>
-
-lemma integrable_on_little_subintervals:
-  fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "\<forall>x\<in>cbox a b. \<exists>d>0. \<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
-    f integrable_on cbox u v"
-  shows "f integrable_on cbox a b"
-proof -
-  have "\<forall>x. \<exists>d. x\<in>cbox a b \<longrightarrow> d>0 \<and> (\<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow>
-    f integrable_on cbox u v)"
-    using assms by auto
-  note this[unfolded gauge_existence_lemma]
-  from choice[OF this] guess d .. note d=this[rule_format]
-  guess p
-    apply (rule fine_division_exists[OF gauge_ball_dependent,of d a b])
-    using d
-    by auto
-  note p=this(1-2)
-  note division_of_tagged_division[OF this(1)]
-  note * = comm_monoid_set.operative_division[OF comm_monoid_set_and operative_integrable, OF this, symmetric, of f]
-  show ?thesis
-    unfolding * comm_monoid_set_F_and
-    apply safe
-    unfolding snd_conv
-  proof -
-    fix x k
-    assume "(x, k) \<in> p"
-    note tagged_division_ofD(2-4)[OF p(1) this] fineD[OF p(2) this]
-    then show "f integrable_on k"
-      apply safe
-      apply (rule d[THEN conjunct2,rule_format,of x])
-      apply (auto intro: order.trans)
-      done
-  qed
-qed
-
-
-subsection \<open>Second FTC or existence of antiderivative.\<close>
-
-lemma integrable_const[intro]: "(\<lambda>x. c) integrable_on cbox a b"
-  unfolding integrable_on_def
-  apply rule
-  apply (rule has_integral_const)
-  done
-
-lemma integral_has_vector_derivative_continuous_at:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes f: "f integrable_on {a..b}"
-      and x: "x \<in> {a..b}"
-      and fx: "continuous (at x within {a..b}) f"
-  shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
-proof -
-  let ?I = "\<lambda>a b. integral {a..b} f"
-  { fix e::real
-    assume "e > 0"
-    obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {a..b}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> norm(f x' - f x) \<le> e"
-      using \<open>e>0\<close> fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le)
-    have "norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
-           if y: "y \<in> {a..b}" and yx: "\<bar>y - x\<bar> < d" for y
-    proof (cases "y < x")
-      case False
-      have "f integrable_on {a..y}"
-        using f y by (simp add: integrable_subinterval_real)
-      then have Idiff: "?I a y - ?I a x = ?I x y"
-        using False x by (simp add: algebra_simps integral_combine)
-      have fux_int: "((\<lambda>u. f u - f x) has_integral integral {x..y} f - (y - x) *\<^sub>R f x) {x..y}"
-        apply (rule has_integral_sub)
-        using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]])
-        using has_integral_const_real [of "f x" x y] False
-        apply (simp add: )
-        done
-      show ?thesis
-        using False
-        apply (simp add: abs_eq_content del: content_real_if)
-        apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
-        using yx False d x y \<open>e>0\<close> apply (auto simp add: Idiff fux_int)
-        done
-    next
-      case True
-      have "f integrable_on {a..x}"
-        using f x by (simp add: integrable_subinterval_real)
-      then have Idiff: "?I a x - ?I a y = ?I y x"
-        using True x y by (simp add: algebra_simps integral_combine)
-      have fux_int: "((\<lambda>u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}"
-        apply (rule has_integral_sub)
-        using x y apply (force intro: integrable_integral [OF integrable_subinterval_real [OF f]])
-        using has_integral_const_real [of "f x" y x] True
-        apply (simp add: )
-        done
-      have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
-        using True
-        apply (simp add: abs_eq_content del: content_real_if)
-        apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"])
-        using yx True d x y \<open>e>0\<close> apply (auto simp add: Idiff fux_int)
-        done
-      then show ?thesis
-        by (simp add: algebra_simps norm_minus_commute)
-    qed
-    then have "\<exists>d>0. \<forall>y\<in>{a..b}. \<bar>y - x\<bar> < d \<longrightarrow> norm (integral {a..y} f - integral {a..x} f - (y - x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
-      using \<open>d>0\<close> by blast
-  }
-  then show ?thesis
-    by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left)
-qed
-
-lemma integral_has_vector_derivative:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a .. b} f"
-    and "x \<in> {a .. b}"
-  shows "((\<lambda>u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b})"
-apply (rule integral_has_vector_derivative_continuous_at [OF integrable_continuous_real])
-using assms
-apply (auto simp: continuous_on_eq_continuous_within)
-done
-
-lemma antiderivative_continuous:
-  fixes q b :: real
-  assumes "continuous_on {a .. b} f"
-  obtains g where "\<forall>x\<in>{a .. b}. (g has_vector_derivative (f x::_::banach)) (at x within {a .. b})"
-  apply (rule that)
-  apply rule
-  using integral_has_vector_derivative[OF assms]
-  apply auto
-  done
-
-
-subsection \<open>Combined fundamental theorem of calculus.\<close>
-
-lemma antiderivative_integral_continuous:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "continuous_on {a .. b} f"
-  obtains g where "\<forall>u\<in>{a .. b}. \<forall>v \<in> {a .. b}. u \<le> v \<longrightarrow> (f has_integral (g v - g u)) {u .. v}"
-proof -
-  from antiderivative_continuous[OF assms] guess g . note g=this
-  show ?thesis
-    apply (rule that[of g])
-    apply safe
-  proof goal_cases
-    case prems: (1 u v)
-    have "\<forall>x\<in>cbox u v. (g has_vector_derivative f x) (at x within cbox u v)"
-      apply rule
-      apply (rule has_vector_derivative_within_subset)
-      apply (rule g[rule_format])
-      using prems(1,2)
-      apply auto
-      done
-    then show ?case
-      using fundamental_theorem_of_calculus[OF prems(3), of g f] by auto
-  qed
-qed
-
-
-subsection \<open>General "twiddling" for interval-to-interval function image.\<close>
-
-lemma has_integral_twiddle:
-  assumes "0 < r"
-    and "\<forall>x. h(g x) = x"
-    and "\<forall>x. g(h x) = x"
-    and "\<forall>x. continuous (at x) g"
-    and "\<forall>u v. \<exists>w z. g ` cbox u v = cbox w z"
-    and "\<forall>u v. \<exists>w z. h ` cbox u v = cbox w z"
-    and "\<forall>u v. content(g ` cbox u v) = r * content (cbox u v)"
-    and "(f has_integral i) (cbox a b)"
-  shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)"
-proof -
-  show ?thesis when *: "cbox a b \<noteq> {} \<Longrightarrow> ?thesis"
-    apply cases
-    defer
-    apply (rule *)
-    apply assumption
-  proof goal_cases
-    case prems: 1
-    then show ?thesis
-      unfolding prems assms(8)[unfolded prems has_integral_empty_eq] by auto
-  qed
-  assume "cbox a b \<noteq> {}"
-  from assms(6)[rule_format,of a b] guess w z by (elim exE) note wz=this
-  have inj: "inj g" "inj h"
-    unfolding inj_on_def
-    apply safe
-    apply(rule_tac[!] ccontr)
-    using assms(2)
-    apply(erule_tac x=x in allE)
-    using assms(2)
-    apply(erule_tac x=y in allE)
-    defer
-    using assms(3)
-    apply (erule_tac x=x in allE)
-    using assms(3)
-    apply(erule_tac x=y in allE)
-    apply auto
-    done
-  show ?thesis
-    unfolding has_integral_def has_integral_compact_interval_def
-    apply (subst if_P)
-    apply rule
-    apply rule
-    apply (rule wz)
-  proof safe
-    fix e :: real
-    assume e: "e > 0"
-    with assms(1) have "e * r > 0" by simp
-    from assms(8)[unfolded has_integral,rule_format,OF this] guess d by (elim exE conjE) note d=this[rule_format]
-    define d' where "d' x = {y. g y \<in> d (g x)}" for x
-    have d': "\<And>x. d' x = {y. g y \<in> (d (g x))}"
-      unfolding d'_def ..
-    show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` cbox a b \<and> d fine p \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)"
-    proof (rule_tac x=d' in exI, safe)
-      show "gauge d'"
-        using d(1)
-        unfolding gauge_def d'
-        using continuous_open_preimage_univ[OF assms(4)]
-        by auto
-      fix p
-      assume as: "p tagged_division_of h ` cbox a b" "d' fine p"
-      note p = tagged_division_ofD[OF as(1)]
-      have "(\<lambda>(x, k). (g x, g ` k)) ` p tagged_division_of (cbox a b) \<and> d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
-        unfolding tagged_division_of
-      proof safe
-        show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)"
-          using as by auto
-        show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p"
-          using as(2) unfolding fine_def d' by auto
-        fix x k
-        assume xk[intro]: "(x, k) \<in> p"
-        show "g x \<in> g ` k"
-          using p(2)[OF xk] by auto
-        show "\<exists>u v. g ` k = cbox u v"
-          using p(4)[OF xk] using assms(5-6) by auto
-        {
-          fix y
-          assume "y \<in> k"
-          then show "g y \<in> cbox a b" "g y \<in> cbox a b"
-            using p(3)[OF xk,unfolded subset_eq,rule_format,of "h (g y)"]
-            using assms(2)[rule_format,of y]
-            unfolding inj_image_mem_iff[OF inj(2)]
-            by auto
-        }
-        fix x' k'
-        assume xk': "(x', k') \<in> p"
-        fix z
-        assume z: "z \<in> interior (g ` k)" "z \<in> interior (g ` k')"
-        have same: "(x, k) = (x', k')"
-          apply -
-          apply (rule ccontr)
-          apply (drule p(5)[OF xk xk'])
-        proof -
-          assume as: "interior k \<inter> interior k' = {}"
-          have "z \<in> g ` (interior k \<inter> interior k')"
-            using interior_image_subset[OF assms(4) inj(1)] z
-            unfolding image_Int[OF inj(1)] by blast
-          then show False
-            using as by blast
-        qed
-        then show "g x = g x'"
-          by auto
-        {
-          fix z
-          assume "z \<in> k"
-          then show "g z \<in> g ` k'"
-            using same by auto
-        }
-        {
-          fix z
-          assume "z \<in> k'"
-          then show "g z \<in> g ` k"
-            using same by auto
-        }
-      next
-        fix x
-        assume "x \<in> cbox a b"
-        then have "h x \<in>  \<Union>{k. \<exists>x. (x, k) \<in> p}"
-          using p(6) by auto
-        then guess X unfolding Union_iff .. note X=this
-        from this(1) guess y unfolding mem_Collect_eq ..
-        then show "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}"
-          apply -
-          apply (rule_tac X="g ` X" in UnionI)
-          defer
-          apply (rule_tac x="h x" in image_eqI)
-          using X(2) assms(3)[rule_format,of x]
-          apply auto
-          done
-      qed
-        note ** = d(2)[OF this]
-        have *: "inj_on (\<lambda>(x, k). (g x, g ` k)) p"
-          using inj(1) unfolding inj_on_def by fastforce
-        have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _")
-          using assms(7)
-          apply (simp only: algebra_simps add_left_cancel scaleR_right.setsum)
-          apply (subst setsum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p])
-          apply (auto intro!: * setsum.cong simp: bij_betw_def dest!: p(4))
-          done
-      also have "\<dots> = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r")
-        unfolding scaleR_diff_right scaleR_scaleR
-        using assms(1)
-        by auto
-      finally have *: "?l = ?r" .
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e"
-        using **
-        unfolding *
-        unfolding norm_scaleR
-        using assms(1)
-        by (auto simp add:field_simps)
-    qed
-  qed
-qed
-
-
-subsection \<open>Special case of a basic affine transformation.\<close>
-
-lemma interval_image_affinity_interval:
-  "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v"
-  unfolding image_affinity_cbox
-  by auto
-
-lemma content_image_affinity_cbox:
-  "content((\<lambda>x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) =
-    \<bar>m\<bar> ^ DIM('a) * content (cbox a b)" (is "?l = ?r")
-proof (cases "cbox a b = {}")
-  case True then show ?thesis by simp
-next
-  case False
-  show ?thesis
-  proof (cases "m \<ge> 0")
-    case True
-    with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \<noteq> {}"
-      unfolding box_ne_empty
-      apply (intro ballI)
-      apply (erule_tac x=i in ballE)
-      apply (auto simp: inner_simps mult_left_mono)
-      done
-    moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b - a) \<bullet> i"
-      by (simp add: inner_simps field_simps)
-    ultimately show ?thesis
-      by (simp add: image_affinity_cbox True content_cbox'
-        setprod.distrib setprod_constant inner_diff_left)
-  next
-    case False
-    with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
-      unfolding box_ne_empty
-      apply (intro ballI)
-      apply (erule_tac x=i in ballE)
-      apply (auto simp: inner_simps mult_left_mono)
-      done
-    moreover from False have *: "\<And>i. (m *\<^sub>R a + c) \<bullet> i - (m *\<^sub>R b + c) \<bullet> i = (-m) *\<^sub>R (b - a) \<bullet> i"
-      by (simp add: inner_simps field_simps)
-    ultimately show ?thesis using False
-      by (simp add: image_affinity_cbox content_cbox'
-        setprod.distrib[symmetric] setprod_constant[symmetric] inner_diff_left)
-  qed
-qed
-
-lemma has_integral_affinity:
-  fixes a :: "'a::euclidean_space"
-  assumes "(f has_integral i) (cbox a b)"
-      and "m \<noteq> 0"
-  shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (\<bar>m\<bar> ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)"
-  apply (rule has_integral_twiddle)
-  using assms
-  apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox)
-  apply (rule zero_less_power)
-  unfolding scaleR_right_distrib
-  apply auto
-  done
-
-lemma integrable_affinity:
-  assumes "f integrable_on cbox a b"
-    and "m \<noteq> 0"
-  shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)"
-  using assms
-  unfolding integrable_on_def
-  apply safe
-  apply (drule has_integral_affinity)
-  apply auto
-  done
-
-lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified]
-
-subsection \<open>Special case of stretching coordinate axes separately.\<close>
-
-lemma content_image_stretch_interval:
-  "content ((\<lambda>x::'a::euclidean_space. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)::'a) ` cbox a b) =
-    \<bar>setprod m Basis\<bar> * content (cbox a b)"
-proof (cases "cbox a b = {}")
-  case True
-  then show ?thesis
-    unfolding content_def image_is_empty image_stretch_interval if_P[OF True] by auto
-next
-  case False
-  then have "(\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b \<noteq> {}"
-    by auto
-  then show ?thesis
-    using False
-    unfolding content_def image_stretch_interval
-    apply -
-    unfolding interval_bounds' if_not_P
-    unfolding abs_setprod setprod.distrib[symmetric]
-    apply (rule setprod.cong)
-    apply (rule refl)
-    unfolding lessThan_iff
-    apply (simp only: inner_setsum_left_Basis)
-  proof -
-    fix i :: 'a
-    assume i: "i \<in> Basis"
-    have "(m i < 0 \<or> m i > 0) \<or> m i = 0"
-      by auto
-    then show "max (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) - min (m i * (a \<bullet> i)) (m i * (b \<bullet> i)) =
-      \<bar>m i\<bar> * (b \<bullet> i - a \<bullet> i)"
-      apply -
-      apply (erule disjE)+
-      unfolding min_def max_def
-      using False[unfolded box_ne_empty,rule_format,of i] i
-      apply (auto simp add:field_simps not_le mult_le_cancel_left_neg mult_le_cancel_left_pos)
-      done
-  qed
-qed
-
-lemma has_integral_stretch:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "(f has_integral i) (cbox a b)"
-    and "\<forall>k\<in>Basis. m k \<noteq> 0"
-  shows "((\<lambda>x. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) has_integral
-    ((1/ \<bar>setprod m Basis\<bar>) *\<^sub>R i)) ((\<lambda>x. (\<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b)"
-  apply (rule has_integral_twiddle[where f=f])
-  unfolding zero_less_abs_iff content_image_stretch_interval
-  unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a]
-  using assms
-proof -
-  show "\<forall>y::'a. continuous (at y) (\<lambda>x. (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k))"
-    apply rule
-    apply (rule linear_continuous_at)
-    unfolding linear_linear
-    unfolding linear_iff inner_simps euclidean_eq_iff[where 'a='a]
-    apply (auto simp add: field_simps)
-    done
-qed auto
-
-lemma integrable_stretch:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
-  assumes "f integrable_on cbox a b"
-    and "\<forall>k\<in>Basis. m k \<noteq> 0"
-  shows "(\<lambda>x::'a. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) integrable_on
-    ((\<lambda>x. \<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k) ` cbox a b)"
-  using assms
-  unfolding integrable_on_def
-  apply -
-  apply (erule exE)
-  apply (drule has_integral_stretch)
-  apply assumption
-  apply auto
-  done
-
-subsection \<open>even more special cases.\<close>
-
-lemma uminus_interval_vector[simp]:
-  fixes a b :: "'a::euclidean_space"
-  shows "uminus ` cbox a b = cbox (-b) (-a)"
-  apply (rule set_eqI)
-  apply rule
-  defer
-  unfolding image_iff
-  apply (rule_tac x="-x" in bexI)
-  apply (auto simp add:minus_le_iff le_minus_iff mem_box)
-  done
-
-lemma has_integral_reflect_lemma[intro]:
-  assumes "(f has_integral i) (cbox a b)"
-  shows "((\<lambda>x. f(-x)) has_integral i) (cbox (-b) (-a))"
-  using has_integral_affinity[OF assms, of "-1" 0]
-  by auto
-
-lemma has_integral_reflect_lemma_real[intro]:
-  assumes "(f has_integral i) {a .. b::real}"
-  shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
-  using assms
-  unfolding box_real[symmetric]
-  by (rule has_integral_reflect_lemma)
-
-lemma has_integral_reflect[simp]:
-  "((\<lambda>x. f (-x)) has_integral i) (cbox (-b) (-a)) \<longleftrightarrow> (f has_integral i) (cbox a b)"
-  apply rule
-  apply (drule_tac[!] has_integral_reflect_lemma)
-  apply auto
-  done
-
-lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on cbox (-b) (-a) \<longleftrightarrow> f integrable_on cbox a b"
-  unfolding integrable_on_def by auto
-
-lemma integrable_reflect_real[simp]: "(\<lambda>x. f(-x)) integrable_on {-b .. -a} \<longleftrightarrow> f integrable_on {a .. b::real}"
-  unfolding box_real[symmetric]
-  by (rule integrable_reflect)
-
-lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\<lambda>x. f (-x)) = integral (cbox a b) f"
-  unfolding integral_def by auto
-
-lemma integral_reflect_real[simp]: "integral {-b .. -a} (\<lambda>x. f (-x)) = integral {a .. b::real} f"
-  unfolding box_real[symmetric]
-  by (rule integral_reflect)
-
-
-subsection \<open>Stronger form of FCT; quite a tedious proof.\<close>
-
-lemma bgauge_existence_lemma: "(\<forall>x\<in>s. \<exists>d::real. 0 < d \<and> q d x) \<longleftrightarrow> (\<forall>x. \<exists>d>0. x\<in>s \<longrightarrow> q d x)"
-  by (meson zero_less_one)
-
-lemma additive_tagged_division_1':
-  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b"
-    and "p tagged_division_of {a..b}"
-  shows "setsum (\<lambda>(x,k). f (Sup k) - f(Inf k)) p = f b - f a"
-  using additive_tagged_division_1[OF _ assms(2), of f]
-  using assms(1)
-  by auto
-
-lemma split_minus[simp]: "(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x"
-  by (simp add: split_def)
-
-lemma norm_triangle_le_sub: "norm x + norm y \<le> e \<Longrightarrow> norm (x - y) \<le> e"
-  apply (subst(asm)(2) norm_minus_cancel[symmetric])
-  apply (drule norm_triangle_le)
-  apply (auto simp add: algebra_simps)
-  done
-
-lemma fundamental_theorem_of_calculus_interior:
-  fixes f :: "real \<Rightarrow> 'a::real_normed_vector"
-  assumes "a \<le> b"
-    and "continuous_on {a .. b} f"
-    and "\<forall>x\<in>{a <..< b}. (f has_vector_derivative f'(x)) (at x)"
-  shows "(f' has_integral (f b - f a)) {a .. b}"
-proof -
-  {
-    presume *: "a < b \<Longrightarrow> ?thesis"
-    show ?thesis
-    proof (cases "a < b")
-      case True
-      then show ?thesis by (rule *)
-    next
-      case False
-      then have "a = b"
-        using assms(1) by auto
-      then have *: "cbox a b = {b}" "f b - f a = 0"
-        by (auto simp add:  order_antisym)
-      show ?thesis
-        unfolding *(2)
-        unfolding content_eq_0
-        using * \<open>a = b\<close>
-        by (auto simp: ex_in_conv)
-    qed
-  }
-  assume ab: "a < b"
-  let ?P = "\<lambda>e. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a .. b} \<and> d fine p \<longrightarrow>
-    norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a .. b})"
-  { presume "\<And>e. e > 0 \<Longrightarrow> ?P e" then show ?thesis unfolding has_integral_factor_content_real by auto }
-  fix e :: real
-  assume e: "e > 0"
-  note assms(3)[unfolded has_vector_derivative_def has_derivative_at_alt ball_conj_distrib]
-  note conjunctD2[OF this]
-  note bounded=this(1) and this(2)
-  from this(2) have "\<forall>x\<in>box a b. \<exists>d>0. \<forall>y. norm (y - x) < d \<longrightarrow>
-    norm (f y - f x - (y - x) *\<^sub>R f' x) \<le> e/2 * norm (y - x)"
-    apply -
-    apply safe
-    apply (erule_tac x=x in ballE)
-    apply (erule_tac x="e/2" in allE)
-    using e
-    apply auto
-    done
-  note this[unfolded bgauge_existence_lemma]
-  from choice[OF this] guess d ..
-  note conjunctD2[OF this[rule_format]]
-  note d = this[rule_format]
-  have "bounded (f ` cbox a b)"
-    apply (rule compact_imp_bounded compact_continuous_image)+
-    using compact_cbox assms
-    apply auto
-    done
-  from this[unfolded bounded_pos] guess B .. note B = this[rule_format]
-
-  have "\<exists>da. 0 < da \<and> (\<forall>c. a \<le> c \<and> {a .. c} \<subseteq> {a .. b} \<and> {a .. c} \<subseteq> ball a da \<longrightarrow>
-    norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b - a)) / 4)"
-  proof -
-    have "a \<in> {a .. b}"
-      using ab by auto
-    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
-    note * = this[unfolded continuous_within Lim_within,rule_format]
-    have "(e * (b - a)) / 8 > 0"
-      using e ab by (auto simp add: field_simps)
-    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
-    have "\<exists>l. 0 < l \<and> norm(l *\<^sub>R f' a) \<le> (e * (b - a)) / 8"
-    proof (cases "f' a = 0")
-      case True
-      thus ?thesis using ab e by auto
-    next
-      case False
-      then show ?thesis
-        apply (rule_tac x="(e * (b - a)) / 8 / norm (f' a)" in exI)
-        using ab e
-        apply (auto simp add: field_simps)
-        done
-    qed
-    then guess l .. note l = conjunctD2[OF this]
-    show ?thesis
-      apply (rule_tac x="min k l" in exI)
-      apply safe
-      unfolding min_less_iff_conj
-      apply rule
-      apply (rule l k)+
-    proof -
-      fix c
-      assume as: "a \<le> c" "{a .. c} \<subseteq> {a .. b}" "{a .. c} \<subseteq> ball a (min k l)"
-      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box]
-      have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)"
-        by (rule norm_triangle_ineq4)
-      also have "\<dots> \<le> e * (b - a) / 8 + e * (b - a) / 8"
-      proof (rule add_mono)
-        have "\<bar>c - a\<bar> \<le> \<bar>l\<bar>"
-          using as' by auto
-        then show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b - a) / 8"
-          apply -
-          apply (rule order_trans[OF _ l(2)])
-          unfolding norm_scaleR
-          apply (rule mult_right_mono)
-          apply auto
-          done
-      next
-        show "norm (f c - f a) \<le> e * (b - a) / 8"
-          apply (rule less_imp_le)
-          apply (cases "a = c")
-          defer
-          apply (rule k(2)[unfolded dist_norm])
-          using as' e ab
-          apply (auto simp add: field_simps)
-          done
-      qed
-      finally show "norm (content {a .. c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b - a) / 4"
-        unfolding content_real[OF as(1)] by auto
-    qed
-  qed
-  then guess da .. note da=conjunctD2[OF this,rule_format]
-
-  have "\<exists>db>0. \<forall>c\<le>b. {c .. b} \<subseteq> {a .. b} \<and> {c .. b} \<subseteq> ball b db \<longrightarrow>
-    norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b - a)) / 4"
-  proof -
-    have "b \<in> {a .. b}"
-      using ab by auto
-    note assms(2)[unfolded continuous_on_eq_continuous_within,rule_format,OF this]
-    note * = this[unfolded continuous_within Lim_within,rule_format] have "(e * (b - a)) / 8 > 0"
-      using e ab by (auto simp add: field_simps)
-    from *[OF this] guess k .. note k = conjunctD2[OF this,rule_format]
-    have "\<exists>l. 0 < l \<and> norm (l *\<^sub>R f' b) \<le> (e * (b - a)) / 8"
-    proof (cases "f' b = 0")
-      case True
-      thus ?thesis using ab e by auto
-    next
-      case False
-      then show ?thesis
-        apply (rule_tac x="(e * (b - a)) / 8 / norm (f' b)" in exI)
-        using ab e
-        apply (auto simp add: field_simps)
-        done
-    qed
-    then guess l .. note l = conjunctD2[OF this]
-    show ?thesis
-      apply (rule_tac x="min k l" in exI)
-      apply safe
-      unfolding min_less_iff_conj
-      apply rule
-      apply (rule l k)+
-    proof -
-      fix c
-      assume as: "c \<le> b" "{c..b} \<subseteq> {a..b}" "{c..b} \<subseteq> ball b (min k l)"
-      note as' = this[unfolded subset_eq Ball_def mem_ball dist_real_def mem_box]
-      have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)"
-        by (rule norm_triangle_ineq4)
-      also have "\<dots> \<le> e * (b - a) / 8 + e * (b - a) / 8"
-      proof (rule add_mono)
-        have "\<bar>c - b\<bar> \<le> \<bar>l\<bar>"
-          using as' by auto
-        then show "norm ((b - c) *\<^sub>R f' b) \<le> e * (b - a) / 8"
-          apply -
-          apply (rule order_trans[OF _ l(2)])
-          unfolding norm_scaleR
-          apply (rule mult_right_mono)
-          apply auto
-          done
-      next
-        show "norm (f b - f c) \<le> e * (b - a) / 8"
-          apply (rule less_imp_le)
-          apply (cases "b = c")
-          defer
-          apply (subst norm_minus_commute)
-          apply (rule k(2)[unfolded dist_norm])
-          using as' e ab
-          apply (auto simp add: field_simps)
-          done
-      qed
-      finally show "norm (content {c .. b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b - a) / 4"
-        unfolding content_real[OF as(1)] by auto
-    qed
-  qed
-  then guess db .. note db=conjunctD2[OF this,rule_format]
-
-  let ?d = "(\<lambda>x. ball x (if x=a then da else if x=b then db else d x))"
-  show "?P e"
-    apply (rule_tac x="?d" in exI)
-  proof (safe, goal_cases)
-    case 1
-    show ?case
-      apply (rule gauge_ball_dependent)
-      using ab db(1) da(1) d(1)
-      apply auto
-      done
-  next
-    case as: (2 p)
-    let ?A = "{t. fst t \<in> {a, b}}"
-    note p = tagged_division_ofD[OF as(1)]
-    have pA: "p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"
-      using as by auto
-    note * = additive_tagged_division_1'[OF assms(1) as(1), symmetric]
-    have **: "\<And>n1 s1 n2 s2::real. n2 \<le> s2 / 2 \<Longrightarrow> n1 - s1 \<le> s2 / 2 \<Longrightarrow> n1 + n2 \<le> s1 + s2"
-      by arith
-    show ?case
-      unfolding content_real[OF assms(1)] and *[of "\<lambda>x. x"] *[of f] setsum_subtractf[symmetric] split_minus
-      unfolding setsum_right_distrib
-      apply (subst(2) pA)
-      apply (subst pA)
-      unfolding setsum.union_disjoint[OF pA(2-)]
-    proof (rule norm_triangle_le, rule **, goal_cases)
-      case 1
-      show ?case
-        apply (rule order_trans)
-        apply (rule setsum_norm_le)
-        defer
-        apply (subst setsum_divide_distrib)
-        apply (rule order_refl)
-        apply safe
-        apply (unfold not_le o_def split_conv fst_conv)
-      proof (rule ccontr)
-        fix x k
-        assume xk: "(x, k) \<in> p"
-          "e * (Sup k -  Inf k) / 2 <
-            norm (content k *\<^sub>R f' x - (f (Sup k) - f (Inf k)))"
-        from p(4)[OF this(1)] guess u v by (elim exE) note k=this
-        then have "u \<le> v" and uv: "{u, v} \<subseteq> cbox u v"
-          using p(2)[OF xk(1)] by auto
-        note result = xk(2)[unfolded k box_real interval_bounds_real[OF this(1)] content_real[OF this(1)]]
-
-        assume as': "x \<noteq> a" "x \<noteq> b"
-        then have "x \<in> box a b"
-          using p(2-3)[OF xk(1)] by (auto simp: mem_box)
-        note  * = d(2)[OF this]
-        have "norm ((v - u) *\<^sub>R f' (x) - (f (v) - f (u))) =
-          norm ((f (u) - f (x) - (u - x) *\<^sub>R f' (x)) - (f (v) - f (x) - (v - x) *\<^sub>R f' (x)))"
-          apply (rule arg_cong[of _ _ norm])
-          unfolding scaleR_left.diff
-          apply auto
-          done
-        also have "\<dots> \<le> e / 2 * norm (u - x) + e / 2 * norm (v - x)"
-          apply (rule norm_triangle_le_sub)
-          apply (rule add_mono)
-          apply (rule_tac[!] *)
-          using fineD[OF as(2) xk(1)] as'
-          unfolding k subset_eq
-          apply -
-          apply (erule_tac x=u in ballE)
-          apply (erule_tac[3] x=v in ballE)
-          using uv
-          apply (auto simp:dist_real_def)
-          done
-        also have "\<dots> \<le> e / 2 * norm (v - u)"
-          using p(2)[OF xk(1)]
-          unfolding k
-          by (auto simp add: field_simps)
-        finally have "e * (v - u) / 2 < e * (v - u) / 2"
-          apply -
-          apply (rule less_le_trans[OF result])
-          using uv
-          apply auto
-          done
-        then show False by auto
-      qed
-    next
-      have *: "\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2) / 2 \<Longrightarrow> x - s1 \<le> s2 / 2"
-        by auto
-      case 2
-      show ?case
-        apply (rule *)
-        apply (rule setsum_nonneg)
-        apply rule
-        apply (unfold split_paired_all split_conv)
-        defer
-        unfolding setsum.union_disjoint[OF pA(2-),symmetric] pA(1)[symmetric]
-        unfolding setsum_right_distrib[symmetric]
-        apply (subst additive_tagged_division_1[OF _ as(1)])
-        apply (rule assms)
-      proof -
-        fix x k
-        assume "(x, k) \<in> p \<inter> {t. fst t \<in> {a, b}}"
-        note xk=IntD1[OF this]
-        from p(4)[OF this] guess u v by (elim exE) note uv=this
-        with p(2)[OF xk] have "cbox u v \<noteq> {}"
-          by auto
-        then show "0 \<le> e * ((Sup k) - (Inf k))"
-          unfolding uv using e by (auto simp add: field_simps)
-      next
-        have *: "\<And>s f t e. setsum f s = setsum f t \<Longrightarrow> norm (setsum f t) \<le> e \<Longrightarrow> norm (setsum f s) \<le> e"
-          by auto
-        show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x -
-          (f ((Sup k)) - f ((Inf k)))) \<le> e * (b - a) / 2"
-          apply (rule *[where t1="p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0}"])
-          apply (rule setsum.mono_neutral_right[OF pA(2)])
-          defer
-          apply rule
-          unfolding split_paired_all split_conv o_def
-        proof goal_cases
-          fix x k
-          assume "(x, k) \<in> p \<inter> {t. fst t \<in> {a, b}} - p \<inter> {t. fst t \<in> {a, b} \<and> content (snd t) \<noteq> 0}"
-          then have xk: "(x, k) \<in> p" "content k = 0"
-            by auto
-          from p(4)[OF xk(1)] guess u v by (elim exE) note uv=this
-          have "k \<noteq> {}"
-            using p(2)[OF xk(1)] by auto
-          then have *: "u = v"
-            using xk
-            unfolding uv content_eq_0 box_eq_empty
-            by auto
-          then show "content k *\<^sub>R (f' (x)) - (f ((Sup k)) - f ((Inf k))) = 0"
-            using xk unfolding uv by auto
-        next
-          have *: "p \<inter> {t. fst t \<in> {a, b} \<and> content(snd t) \<noteq> 0} =
-            {t. t\<in>p \<and> fst t = a \<and> content(snd t) \<noteq> 0} \<union> {t. t\<in>p \<and> fst t = b \<and> content(snd t) \<noteq> 0}"
-            by blast
-          have **: "norm (setsum f s) \<le> e"
-            if "\<forall>x y. x \<in> s \<and> y \<in> s \<longrightarrow> x = y"
-            and "\<forall>x. x \<in> s \<longrightarrow> norm (f x) \<le> e"
-            and "e > 0"
-            for s f and e :: real
-          proof (cases "s = {}")
-            case True
-            with that show ?thesis by auto
-          next
-            case False
-            then obtain x where "x \<in> s"
-              by auto
-            then have *: "s = {x}"
-              using that(1) by auto
-            then show ?thesis
-              using \<open>x \<in> s\<close> that(2) by auto
-          qed
-          case 2
-          show ?case
-            apply (subst *)
-            apply (subst setsum.union_disjoint)
-            prefer 4
-            apply (rule order_trans[of _ "e * (b - a)/4 + e * (b - a)/4"])
-            apply (rule norm_triangle_le,rule add_mono)
-            apply (rule_tac[1-2] **)
-          proof -
-            let ?B = "\<lambda>x. {t \<in> p. fst t = x \<and> content (snd t) \<noteq> 0}"
-            have pa: "\<exists>v. k = cbox a v \<and> a \<le> v" if "(a, k) \<in> p" for k
-            proof -
-              guess u v using p(4)[OF that] by (elim exE) note uv=this
-              have *: "u \<le> v"
-                using p(2)[OF that] unfolding uv by auto
-              have u: "u = a"
-              proof (rule ccontr)
-                have "u \<in> cbox u v"
-                  using p(2-3)[OF that(1)] unfolding uv by auto
-                have "u \<ge> a"
-                  using p(2-3)[OF that(1)] unfolding uv subset_eq by auto
-                moreover assume "\<not> ?thesis"
-                ultimately have "u > a" by auto
-                then show False
-                  using p(2)[OF that(1)] unfolding uv by (auto simp add:)
-              qed
-              then show ?thesis
-                apply (rule_tac x=v in exI)
-                unfolding uv
-                using *
-                apply auto
-                done
-            qed
-            have pb: "\<exists>v. k = cbox v b \<and> b \<ge> v" if "(b, k) \<in> p" for k
-            proof -
-              guess u v using p(4)[OF that] by (elim exE) note uv=this
-              have *: "u \<le> v"
-                using p(2)[OF that] unfolding uv by auto
-              have u: "v = b"
-              proof (rule ccontr)
-                have "u \<in> cbox u v"
-                  using p(2-3)[OF that(1)] unfolding uv by auto
-                have "v \<le> b"
-                  using p(2-3)[OF that(1)] unfolding uv subset_eq by auto
-                moreover assume "\<not> ?thesis"
-                ultimately have "v < b" by auto
-                then show False
-                  using p(2)[OF that(1)] unfolding uv by (auto simp add:)
-              qed
-              then show ?thesis
-                apply (rule_tac x=u in exI)
-                unfolding uv
-                using *
-                apply auto
-                done
-            qed
-            show "\<forall>x y. x \<in> ?B a \<and> y \<in> ?B a \<longrightarrow> x = y"
-              apply (rule,rule,rule,unfold split_paired_all)
-              unfolding mem_Collect_eq fst_conv snd_conv
-              apply safe
-            proof -
-              fix x k k'
-              assume k: "(a, k) \<in> p" "(a, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
-              guess v using pa[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pa[OF k(2)] .. note v' = conjunctD2[OF this] let ?v = "min v v'"
-              have "box a ?v \<subseteq> k \<inter> k'"
-                unfolding v v' by (auto simp add: mem_box)
-              note interior_mono[OF this,unfolded interior_Int]
-              moreover have "(a + ?v)/2 \<in> box a ?v"
-                using k(3-)
-                unfolding v v' content_eq_0 not_le
-                by (auto simp add: mem_box)
-              ultimately have "(a + ?v)/2 \<in> interior k \<inter> interior k'"
-                unfolding interior_open[OF open_box] by auto
-              then have *: "k = k'"
-                apply -
-                apply (rule ccontr)
-                using p(5)[OF k(1-2)]
-                apply auto
-                done
-              { assume "x \<in> k" then show "x \<in> k'" unfolding * . }
-              { assume "x \<in> k'" then show "x \<in> k" unfolding * . }
-            qed
-            show "\<forall>x y. x \<in> ?B b \<and> y \<in> ?B b \<longrightarrow> x = y"
-              apply rule
-              apply rule
-              apply rule
-              apply (unfold split_paired_all)
-              unfolding mem_Collect_eq fst_conv snd_conv
-              apply safe
-            proof -
-              fix x k k'
-              assume k: "(b, k) \<in> p" "(b, k') \<in> p" "content k \<noteq> 0" "content k' \<noteq> 0"
-              guess v using pb[OF k(1)] .. note v = conjunctD2[OF this]
-              guess v' using pb[OF k(2)] .. note v' = conjunctD2[OF this]
-              let ?v = "max v v'"
-              have "box ?v b \<subseteq> k \<inter> k'"
-                unfolding v v' by (auto simp: mem_box)
-                note interior_mono[OF this,unfolded interior_Int]
-              moreover have " ((b + ?v)/2) \<in> box ?v b"
-                using k(3-) unfolding v v' content_eq_0 not_le by (auto simp: mem_box)
-              ultimately have " ((b + ?v)/2) \<in> interior k \<inter> interior k'"
-                unfolding interior_open[OF open_box] by auto
-              then have *: "k = k'"
-                apply -
-                apply (rule ccontr)
-                using p(5)[OF k(1-2)]
-                apply auto
-                done
-              { assume "x \<in> k" then show "x \<in> k'" unfolding * . }
-              { assume "x \<in> k'" then show "x\<in>k" unfolding * . }
-            qed
-
-            let ?a = a and ?b = b (* a is something else while proofing the next theorem. *)
-            show "\<forall>x. x \<in> ?B a \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' x - (f (Sup k) -
-              f (Inf k))) x) \<le> e * (b - a) / 4"
-              apply rule
-              apply rule
-              unfolding mem_Collect_eq
-              unfolding split_paired_all fst_conv snd_conv
-            proof (safe, goal_cases)
-              case prems: 1
-              guess v using pa[OF prems(1)] .. note v = conjunctD2[OF this]
-              have "?a \<in> {?a..v}"
-                using v(2) by auto
-              then have "v \<le> ?b"
-                using p(3)[OF prems(1)] unfolding subset_eq v by auto
-              moreover have "{?a..v} \<subseteq> ball ?a da"
-                using fineD[OF as(2) prems(1)]
-                apply -
-                apply (subst(asm) if_P)
-                apply (rule refl)
-                unfolding subset_eq
-                apply safe
-                apply (erule_tac x=" x" in ballE)
-                apply (auto simp add:subset_eq dist_real_def v)
-                done
-              ultimately show ?case
-                unfolding v interval_bounds_real[OF v(2)] box_real
-                apply -
-                apply(rule da(2)[of "v"])
-                using prems fineD[OF as(2) prems(1)]
-                unfolding v content_eq_0
-                apply auto
-                done
-            qed
-            show "\<forall>x. x \<in> ?B b \<longrightarrow> norm ((\<lambda>(x, k). content k *\<^sub>R f' x -
-              (f (Sup k) - f (Inf k))) x) \<le> e * (b - a) / 4"
-              apply rule
-              apply rule
-              unfolding mem_Collect_eq
-              unfolding split_paired_all fst_conv snd_conv
-            proof (safe, goal_cases)
-              case prems: 1
-              guess v using pb[OF prems(1)] .. note v = conjunctD2[OF this]
-              have "?b \<in> {v.. ?b}"
-                using v(2) by auto
-              then have "v \<ge> ?a" using p(3)[OF prems(1)]
-                unfolding subset_eq v by auto
-              moreover have "{v..?b} \<subseteq> ball ?b db"
-                using fineD[OF as(2) prems(1)]
-                apply -
-                apply (subst(asm) if_P, rule refl)
-                unfolding subset_eq
-                apply safe
-                apply (erule_tac x=" x" in ballE)
-                using ab
-                apply (auto simp add:subset_eq v dist_real_def)
-                done
-              ultimately show ?case
-                unfolding v
-                unfolding interval_bounds_real[OF v(2)] box_real
-                apply -
-                apply(rule db(2)[of "v"])
-                using prems fineD[OF as(2) prems(1)]
-                unfolding v content_eq_0
-                apply auto
-                done
-            qed
-          qed (insert p(1) ab e, auto simp add: field_simps)
-        qed auto
-      qed
-    qed
-  qed
-qed
-
-
-subsection \<open>Stronger form with finite number of exceptional points.\<close>
-
-lemma fundamental_theorem_of_calculus_interior_strong:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "finite s"
-    and "a \<le> b"
-    and "continuous_on {a .. b} f"
-    and "\<forall>x\<in>{a <..< b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "(f' has_integral (f b - f a)) {a .. b}"
-  using assms
-proof (induct "card s" arbitrary: s a b)
-  case 0
-  show ?case
-    apply (rule fundamental_theorem_of_calculus_interior)
-    using 0
-    apply auto
-    done
-next
-  case (Suc n)
-  from this(2) guess c s'
-    apply -
-    apply (subst(asm) eq_commute)
-    unfolding card_Suc_eq
-    apply (subst(asm)(2) eq_commute)
-    apply (elim exE conjE)
-    done
-  note cs = this[rule_format]
-  show ?case
-  proof (cases "c \<in> box a b")
-    case False
-    then show ?thesis
-      apply -
-      apply (rule Suc(1)[OF cs(3) _ Suc(4,5)])
-      apply safe
-      defer
-      apply (rule Suc(6)[rule_format])
-      using Suc(3)
-      unfolding cs
-      apply auto
-      done
-  next
-    have *: "f b - f a = (f c - f a) + (f b - f c)"
-      by auto
-    case True
-    then have "a \<le> c" "c \<le> b"
-      by (auto simp: mem_box)
-    then show ?thesis
-      apply (subst *)
-      apply (rule has_integral_combine)
-      apply assumption+
-      apply (rule_tac[!] Suc(1)[OF cs(3)])
-      using Suc(3)
-      unfolding cs
-    proof -
-      show "continuous_on {a .. c} f" "continuous_on {c .. b} f"
-        apply (rule_tac[!] continuous_on_subset[OF Suc(5)])
-        using True
-        apply (auto simp: mem_box)
-        done
-      let ?P = "\<lambda>i j. \<forall>x\<in>{i <..< j} - s'. (f has_vector_derivative f' x) (at x)"
-      show "?P a c" "?P c b"
-        apply safe
-        apply (rule_tac[!] Suc(6)[rule_format])
-        using True
-        unfolding cs
-        apply (auto simp: mem_box)
-        done
-    qed auto
-  qed
-qed
-
-lemma fundamental_theorem_of_calculus_strong:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "finite s"
-    and "a \<le> b"
-    and "continuous_on {a .. b} f"
-    and "\<forall>x\<in>{a .. b} - s. (f has_vector_derivative f'(x)) (at x)"
-  shows "(f' has_integral (f b - f a)) {a .. b}"
-  apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f'])
-  using assms(4)
-  apply (auto simp: mem_box)
-  done
-
-lemma indefinite_integral_continuous_left:
-  fixes f:: "real \<Rightarrow> 'a::banach"
-  assumes "f integrable_on {a .. b}"
-    and "a < c"
-    and "c \<le> b"
-    and "e > 0"
-  obtains d where "d > 0"
-    and "\<forall>t. c - d < t \<and> t \<le> c \<longrightarrow> norm (integral {a .. c} f - integral {a .. t} f) < e"
-proof -
-  have "\<exists>w>0. \<forall>t. c - w < t \<and> t < c \<longrightarrow> norm (f c) * norm(c - t) < e / 3"
-  proof (cases "f c = 0")
-    case False
-    hence "0 < e / 3 / norm (f c)" using \<open>e>0\<close> by simp
-    then show ?thesis
-      apply -
-      apply rule
-      apply rule
-      apply assumption
-      apply safe
-    proof -
-      fix t
-      assume as: "t < c" and "c - e / 3 / norm (f c) < t"
-      then have "c - t < e / 3 / norm (f c)"
-        by auto
-      then have "norm (c - t) < e / 3 / norm (f c)"
-        using as by auto
-      then show "norm (f c) * norm (c - t) < e / 3"
-        using False
-        apply -
-        apply (subst mult.commute)
-        apply (subst pos_less_divide_eq[symmetric])
-        apply auto
-        done
-    qed
-  next
-    case True
-    show ?thesis
-      apply (rule_tac x=1 in exI)
-      unfolding True
-      using \<open>e > 0\<close>
-      apply auto
-      done
-  qed
-  then guess w .. note w = conjunctD2[OF this,rule_format]
-
-  have *: "e / 3 > 0"
-    using assms by auto
-  have "f integrable_on {a .. c}"
-    apply (rule integrable_subinterval_real[OF assms(1)])
-    using assms(2-3)
-    apply auto
-    done
-  from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d1 ..
-  note d1 = conjunctD2[OF this,rule_format]
-  define d where [abs_def]: "d x = ball x w \<inter> d1 x" for x
-  have "gauge d"
-    unfolding d_def using w(1) d1 by auto
-  note this[unfolded gauge_def,rule_format,of c]
-  note conjunctD2[OF this]
-  from this(2)[unfolded open_contains_ball,rule_format,OF this(1)] guess k ..
-  note k=conjunctD2[OF this]
-
-  let ?d = "min k (c - a) / 2"
-  show ?thesis
-    apply (rule that[of ?d])
-    apply safe
-  proof -
-    show "?d > 0"
-      using k(1) using assms(2) by auto
-    fix t
-    assume as: "c - ?d < t" "t \<le> c"
-    let ?thesis = "norm (integral ({a .. c}) f - integral ({a .. t}) f) < e"
-    {
-      presume *: "t < c \<Longrightarrow> ?thesis"
-      show ?thesis
-        apply (cases "t = c")
-        defer
-        apply (rule *)
-        apply (subst less_le)
-        using \<open>e > 0\<close> as(2)
-        apply auto
-        done
-    }
-    assume "t < c"
-
-    have "f integrable_on {a .. t}"
-      apply (rule integrable_subinterval_real[OF assms(1)])
-      using assms(2-3) as(2)
-      apply auto
-      done
-    from integrable_integral[OF this,unfolded has_integral_real,rule_format,OF *] guess d2 ..
-    note d2 = conjunctD2[OF this,rule_format]
-    define d3 where "d3 x = (if x \<le> t then d1 x \<inter> d2 x else d1 x)" for x
-    have "gauge d3"
-      using d2(1) d1(1) unfolding d3_def gauge_def by auto
-    from fine_division_exists_real[OF this, of a t] guess p . note p=this
-    note p'=tagged_division_ofD[OF this(1)]
-    have pt: "\<forall>(x,k)\<in>p. x \<le> t"
-    proof (safe, goal_cases)
-      case prems: 1
-      from p'(2,3)[OF prems] show ?case
-        by auto
-    qed
-    with p(2) have "d2 fine p"
-      unfolding fine_def d3_def
-      apply safe
-      apply (erule_tac x="(a,b)" in ballE)+
-      apply auto
-      done
-    note d2_fin = d2(2)[OF conjI[OF p(1) this]]
-
-    have *: "{a .. c} \<inter> {x. x \<bullet> 1 \<le> t} = {a .. t}" "{a .. c} \<inter> {x. x \<bullet> 1 \<ge> t} = {t .. c}"
-      using assms(2-3) as by (auto simp add: field_simps)
-    have "p \<union> {(c, {t .. c})} tagged_division_of {a .. c} \<and> d1 fine p \<union> {(c, {t .. c})}"
-      apply rule
-      apply (rule tagged_division_union_interval_real[of _ _ _ 1 "t"])
-      unfolding *
-      apply (rule p)
-      apply (rule tagged_division_of_self_real)
-      unfolding fine_def
-      apply safe
-    proof -
-      fix x k y
-      assume "(x,k) \<in> p" and "y \<in> k"
-      then show "y \<in> d1 x"
-        using p(2) pt
-        unfolding fine_def d3_def
-        apply -
-        apply (erule_tac x="(x,k)" in ballE)+
-        apply auto
-        done
-    next
-      fix x assume "x \<in> {t..c}"
-      then have "dist c x < k"
-        unfolding dist_real_def
-        using as(1)
-        by (auto simp add: field_simps)
-      then show "x \<in> d1 c"
-        using k(2)
-        unfolding d_def
-        by auto
-    qed (insert as(2), auto) note d1_fin = d1(2)[OF this]
-
-    have *: "integral {a .. c} f - integral {a .. t} f = -(((c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)) -
-      integral {a .. c} f) + ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - integral {a .. t} f) + (c - t) *\<^sub>R f c"
-      "e = (e/3 + e/3) + e/3"
-      by auto
-    have **: "(\<Sum>(x, k)\<in>p \<union> {(c, {t .. c})}. content k *\<^sub>R f x) =
-      (c - t) *\<^sub>R f c + (\<Sum>(x, k)\<in>p. content k *\<^sub>R f x)"
-    proof -
-      have **: "\<And>x F. F \<union> {x} = insert x F"
-        by auto
-      have "(c, cbox t c) \<notin> p"
-      proof (safe, goal_cases)
-        case prems: 1
-        from p'(2-3)[OF prems] have "c \<in> cbox a t"
-          by auto
-        then show False using \<open>t < c\<close>
-          by auto
-      qed
-      then show ?thesis
-        unfolding ** box_real
-        apply -
-        apply (subst setsum.insert)
-        apply (rule p')
-        unfolding split_conv
-        defer
-        apply (subst content_real)
-        using as(2)
-        apply auto
-        done
-    qed
-    have ***: "c - w < t \<and> t < c"
-    proof -
-      have "c - k < t"
-        using \<open>k>0\<close> as(1) by (auto simp add: field_simps)
-      moreover have "k \<le> w"
-        apply (rule ccontr)
-        using k(2)
-        unfolding subset_eq
-        apply (erule_tac x="c + ((k + w)/2)" in ballE)
-        unfolding d_def
-        using \<open>k > 0\<close> \<open>w > 0\<close>
-        apply (auto simp add: field_simps not_le not_less dist_real_def)
-        done
-      ultimately show ?thesis using \<open>t < c\<close>
-        by (auto simp add: field_simps)
-    qed
-    show ?thesis
-      unfolding *(1)
-      apply (subst *(2))
-      apply (rule norm_triangle_lt add_strict_mono)+
-      unfolding norm_minus_cancel
-      apply (rule d1_fin[unfolded **])
-      apply (rule d2_fin)
-      using w(2)[OF ***]
-      unfolding norm_scaleR
-      apply (auto simp add: field_simps)
-      done
-  qed
-qed
-
-lemma indefinite_integral_continuous_right:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "f integrable_on {a .. b}"
-    and "a \<le> c"
-    and "c < b"
-    and "e > 0"
-  obtains d where "0 < d"
-    and "\<forall>t. c \<le> t \<and> t < c + d \<longrightarrow> norm (integral {a .. c} f - integral {a .. t} f) < e"
-proof -
-  have *: "(\<lambda>x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \<le> - a"
-    using assms by auto
-  from indefinite_integral_continuous_left[OF * \<open>e>0\<close>] guess d . note d = this
-  let ?d = "min d (b - c)"
-  show ?thesis
-    apply (rule that[of "?d"])
-    apply safe
-  proof -
-    show "0 < ?d"
-      using d(1) assms(3) by auto
-    fix t :: real
-    assume as: "c \<le> t" "t < c + ?d"
-    have *: "integral {a .. c} f = integral {a .. b} f - integral {c .. b} f"
-      "integral {a .. t} f = integral {a .. b} f - integral {t .. b} f"
-      apply (simp_all only: algebra_simps)
-      apply (rule_tac[!] integral_combine)
-      using assms as
-      apply auto
-      done
-    have "(- c) - d < (- t) \<and> - t \<le> - c"
-      using as by auto note d(2)[rule_format,OF this]
-    then show "norm (integral {a .. c} f - integral {a .. t} f) < e"
-      unfolding *
-      unfolding integral_reflect
-      apply (subst norm_minus_commute)
-      apply (auto simp add: algebra_simps)
-      done
-  qed
-qed
-
-lemma indefinite_integral_continuous:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "f integrable_on {a .. b}"
-  shows "continuous_on {a .. b} (\<lambda>x. integral {a .. x} f)"
-proof (unfold continuous_on_iff, safe)
-  fix x e :: real
-  assume as: "x \<in> {a .. b}" "e > 0"
-  let ?thesis = "\<exists>d>0. \<forall>x'\<in>{a .. b}. dist x' x < d \<longrightarrow> dist (integral {a .. x'} f) (integral {a .. x} f) < e"
-  {
-    presume *: "a < b \<Longrightarrow> ?thesis"
-    show ?thesis
-      apply cases
-      apply (rule *)
-      apply assumption
-    proof goal_cases
-      case 1
-      then have "cbox a b = {x}"
-        using as(1)
-        apply -
-        apply (rule set_eqI)
-        apply auto
-        done
-      then show ?case using \<open>e > 0\<close> by auto
-    qed
-  }
-  assume "a < b"
-  have "(x = a \<or> x = b) \<or> (a < x \<and> x < b)"
-    using as(1) by auto
-  then show ?thesis
-    apply (elim disjE)
-  proof -
-    assume "x = a"
-    have "a \<le> a" ..
-    from indefinite_integral_continuous_right[OF assms(1) this \<open>a<b\<close> \<open>e>0\<close>] guess d . note d=this
-    show ?thesis
-      apply rule
-      apply rule
-      apply (rule d)
-      apply safe
-      apply (subst dist_commute)
-      unfolding \<open>x = a\<close> dist_norm
-      apply (rule d(2)[rule_format])
-      apply auto
-      done
-  next
-    assume "x = b"
-    have "b \<le> b" ..
-    from indefinite_integral_continuous_left[OF assms(1) \<open>a<b\<close> this \<open>e>0\<close>] guess d . note d=this
-    show ?thesis
-      apply rule
-      apply rule
-      apply (rule d)
-      apply safe
-      apply (subst dist_commute)
-      unfolding \<open>x = b\<close> dist_norm
-      apply (rule d(2)[rule_format])
-      apply auto
-      done
-  next
-    assume "a < x \<and> x < b"
-    then have xl: "a < x" "x \<le> b" and xr: "a \<le> x" "x < b"
-      by auto
-    from indefinite_integral_continuous_left [OF assms(1) xl \<open>e>0\<close>] guess d1 . note d1=this
-    from indefinite_integral_continuous_right[OF assms(1) xr \<open>e>0\<close>] guess d2 . note d2=this
-    show ?thesis
-      apply (rule_tac x="min d1 d2" in exI)
-    proof safe
-      show "0 < min d1 d2"
-        using d1 d2 by auto
-      fix y
-      assume "y \<in> {a .. b}" and "dist y x < min d1 d2"
-      then show "dist (integral {a .. y} f) (integral {a .. x} f) < e"
-        apply (subst dist_commute)
-        apply (cases "y < x")
-        unfolding dist_norm
-        apply (rule d1(2)[rule_format])
-        defer
-        apply (rule d2(2)[rule_format])
-        unfolding not_less
-        apply (auto simp add: field_simps)
-        done
-    qed
-  qed
-qed
-
-
-subsection \<open>This doesn't directly involve integration, but that gives an easy proof.\<close>
-
-lemma has_derivative_zero_unique_strong_interval:
-  fixes f :: "real \<Rightarrow> 'a::banach"
-  assumes "finite k"
-    and "continuous_on {a .. b} f"
-    and "f a = y"
-    and "\<forall>x\<in>({a .. b} - k). (f has_derivative (\<lambda>h. 0)) (at x within {a .. b})" "x \<in> {a .. b}"
-  shows "f x = y"
-proof -
-  have ab: "a \<le> b"
-    using assms by auto
-  have *: "a \<le> x"
-    using assms(5) by auto
-  have "((\<lambda>x. 0::'a) has_integral f x - f a) {a .. x}"
-    apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *])
-    apply (rule continuous_on_subset[OF assms(2)])
-    defer
-    apply safe
-    unfolding has_vector_derivative_def
-    apply (subst has_derivative_within_open[symmetric])
-    apply assumption
-    apply (rule open_greaterThanLessThan)
-    apply (rule has_derivative_within_subset[where s="{a .. b}"])
-    using assms(4) assms(5)
-    apply (auto simp: mem_box)
-    done
-  note this[unfolded *]
-  note has_integral_unique[OF has_integral_0 this]
-  then show ?thesis
-    unfolding assms by auto
-qed
-
-
-subsection \<open>Generalize a bit to any convex set.\<close>
-
-lemma has_derivative_zero_unique_strong_convex:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "convex s"
-    and "finite k"
-    and "continuous_on s f"
-    and "c \<in> s"
-    and "f c = y"
-    and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
-    and "x \<in> s"
-  shows "f x = y"
-proof -
-  {
-    presume *: "x \<noteq> c \<Longrightarrow> ?thesis"
-    show ?thesis
-      apply cases
-      apply (rule *)
-      apply assumption
-      unfolding assms(5)[symmetric]
-      apply auto
-      done
-  }
-  assume "x \<noteq> c"
-  note conv = assms(1)[unfolded convex_alt,rule_format]
-  have as1: "continuous_on {0 ..1} (f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x))"
-    apply (rule continuous_intros)+
-    apply (rule continuous_on_subset[OF assms(3)])
-    apply safe
-    apply (rule conv)
-    using assms(4,7)
-    apply auto
-    done
-  have *: "t = xa" if "(1 - t) *\<^sub>R c + t *\<^sub>R x = (1 - xa) *\<^sub>R c + xa *\<^sub>R x" for t xa
-  proof -
-    from that have "(t - xa) *\<^sub>R x = (t - xa) *\<^sub>R c"
-      unfolding scaleR_simps by (auto simp add: algebra_simps)
-    then show ?thesis
-      using \<open>x \<noteq> c\<close> by auto
-  qed
-  have as2: "finite {t. ((1 - t) *\<^sub>R c + t *\<^sub>R x) \<in> k}"
-    using assms(2)
-    apply (rule finite_surj[where f="\<lambda>z. SOME t. (1-t) *\<^sub>R c + t *\<^sub>R x = z"])
-    apply safe
-    unfolding image_iff
-    apply rule
-    defer
-    apply assumption
-    apply (rule sym)
-    apply (rule some_equality)
-    defer
-    apply (drule *)
-    apply auto
-    done
-  have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x)) 1 = y"
-    apply (rule has_derivative_zero_unique_strong_interval[OF as2 as1, of ])
-    unfolding o_def
-    using assms(5)
-    defer
-    apply -
-    apply rule
-  proof -
-    fix t
-    assume as: "t \<in> {0 .. 1} - {t. (1 - t) *\<^sub>R c + t *\<^sub>R x \<in> k}"
-    have *: "c - t *\<^sub>R c + t *\<^sub>R x \<in> s - k"
-      apply safe
-      apply (rule conv[unfolded scaleR_simps])
-      using \<open>x \<in> s\<close> \<open>c \<in> s\<close> as
-      by (auto simp add: algebra_simps)
-    have "(f \<circ> (\<lambda>t. (1 - t) *\<^sub>R c + t *\<^sub>R x) has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x))
-      (at t within {0 .. 1})"
-      apply (intro derivative_eq_intros)
-      apply simp_all
-      apply (simp add: field_simps)
-      unfolding scaleR_simps
-      apply (rule has_derivative_within_subset,rule assms(6)[rule_format])
-      apply (rule *)
-      apply safe
-      apply (rule conv[unfolded scaleR_simps])
-      using \<open>x \<in> s\<close> \<open>c \<in> s\<close>
-      apply auto
-      done
-    then show "((\<lambda>xa. f ((1 - xa) *\<^sub>R c + xa *\<^sub>R x)) has_derivative (\<lambda>h. 0)) (at t within {0 .. 1})"
-      unfolding o_def .
-  qed auto
-  then show ?thesis
-    by auto
-qed
-
-
-text \<open>Also to any open connected set with finite set of exceptions. Could
- generalize to locally convex set with limpt-free set of exceptions.\<close>
-
-lemma has_derivative_zero_unique_strong_connected:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "connected s"
-    and "open s"
-    and "finite k"
-    and "continuous_on s f"
-    and "c \<in> s"
-    and "f c = y"
-    and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
-    and "x\<in>s"
-  shows "f x = y"
-proof -
-  have "{x \<in> s. f x \<in> {y}} = {} \<or> {x \<in> s. f x \<in> {y}} = s"
-    apply (rule assms(1)[unfolded connected_clopen,rule_format])
-    apply rule
-    defer
-    apply (rule continuous_closedin_preimage[OF assms(4) closed_singleton])
-    apply (rule open_openin_trans[OF assms(2)])
-    unfolding open_contains_ball
-  proof safe
-    fix x
-    assume "x \<in> s"
-    from assms(2)[unfolded open_contains_ball,rule_format,OF this] guess e .. note e=conjunctD2[OF this]
-    show "\<exists>e>0. ball x e \<subseteq> {xa \<in> s. f xa \<in> {f x}}"
-      apply rule
-      apply rule
-      apply (rule e)
-    proof safe
-      fix y
-      assume y: "y \<in> ball x e"
-      then show "y \<in> s"
-        using e by auto
-      show "f y = f x"
-        apply (rule has_derivative_zero_unique_strong_convex[OF convex_ball])
-        apply (rule assms)
-        apply (rule continuous_on_subset)
-        apply (rule assms)
-        apply (rule e)+
-        apply (subst centre_in_ball)
-        apply (rule e)
-        apply rule
-        apply safe
-        apply (rule has_derivative_within_subset)
-        apply (rule assms(7)[rule_format])
-        using y e
-        apply auto
-        done
-    qed
-  qed
-  then show ?thesis
-    using \<open>x \<in> s\<close> \<open>f c = y\<close> \<open>c \<in> s\<close> by auto
-qed
-
-lemma has_derivative_zero_connected_constant:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "connected s"
-      and "open s"
-      and "finite k"
-      and "continuous_on s f"
-      and "\<forall>x\<in>(s - k). (f has_derivative (\<lambda>h. 0)) (at x within s)"
-    obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c"
-proof (cases "s = {}")
-  case True
-  then show ?thesis
-by (metis empty_iff that)
-next
-  case False
-  then obtain c where "c \<in> s"
-    by (metis equals0I)
-  then show ?thesis
-    by (metis has_derivative_zero_unique_strong_connected assms that)
-qed
-
-
-subsection \<open>Integrating characteristic function of an interval\<close>
-
-lemma has_integral_restrict_open_subinterval:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "(f has_integral i) (cbox c d)"
-    and "cbox c d \<subseteq> cbox a b"
-  shows "((\<lambda>x. if x \<in> box c d then f x else 0) has_integral i) (cbox a b)"
-proof -
-  define g where [abs_def]: "g x = (if x \<in>box c d then f x else 0)" for x
-  {
-    presume *: "cbox c d \<noteq> {} \<Longrightarrow> ?thesis"
-    show ?thesis
-      apply cases
-      apply (rule *)
-      apply assumption
-    proof goal_cases
-      case prems: 1
-      then have *: "box c d = {}"
-        by (metis bot.extremum_uniqueI box_subset_cbox)
-      show ?thesis
-        using assms(1)
-        unfolding *
-        using prems
-        by auto
-    qed
-  }
-  assume "cbox c d \<noteq> {}"
-  from partial_division_extend_1[OF assms(2) this] guess p . note p=this
-  note operat = comm_monoid_set.operative_division[OF add.comm_monoid_set_lift_option operative_integral p(1), symmetric]
-  let ?P = "(if g integrable_on cbox a b then Some (integral (cbox a b) g) else None) = Some i"
-  {
-    presume "?P"
-    then have "g integrable_on cbox a b \<and> integral (cbox a b) g = i"
-      apply -
-      apply cases
-      apply (subst(asm) if_P)
-      apply assumption
-      apply auto
-      done
-    then show ?thesis
-      using integrable_integral
-      unfolding g_def
-      by auto
-  }
-  let ?F = "comm_monoid_set.F (lift_option op +) (Some 0)"
-  have iterate:"?F (\<lambda>i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0"
-  proof (intro comm_monoid_set.neutral[OF add.comm_monoid_set_lift_option] ballI)
-    fix x
-    assume x: "x \<in> p - {cbox c d}"
-    then have "x \<in> p"
-      by auto
-    note div = division_ofD(2-5)[OF p(1) this]
-    from div(3) guess u v by (elim exE) note uv=this
-    have "interior x \<inter> interior (cbox c d) = {}"
-      using div(4)[OF p(2)] x by auto
-    then have "(g has_integral 0) x"
-      unfolding uv
-      apply -
-      apply (rule has_integral_spike_interior[where f="\<lambda>x. 0"])
-      unfolding g_def interior_cbox
-      apply auto
-      done
-    then show "(if g integrable_on x then Some (integral x g) else None) = Some 0"
-      by auto
-  qed
-
-  have *: "p = insert (cbox c d) (p - {cbox c d})"
-    using p by auto
-  have **: "g integrable_on cbox c d"
-    apply (rule integrable_spike_interior[where f=f])
-    unfolding g_def  using assms(1)
-    apply auto
-    done
-  moreover
-  have "integral (cbox c d) g = i"
-    apply (rule has_integral_unique[OF _ assms(1)])
-    apply (rule has_integral_spike_interior[where f=g])
-    defer
-    apply (rule integrable_integral[OF **])
-    unfolding g_def
-    apply auto
-    done
-  ultimately show ?P
-    unfolding operat
-    using p
-    apply (subst *)
-    apply (subst comm_monoid_set.insert[OF add.comm_monoid_set_lift_option])
-    apply (simp_all add: division_of_finite iterate)
-    done
-qed
-
-lemma has_integral_restrict_closed_subinterval:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "(f has_integral i) (cbox c d)"
-    and "cbox c d \<subseteq> cbox a b"
-  shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)"
-proof -
-  note has_integral_restrict_open_subinterval[OF assms]
-  note * = has_integral_spike[OF negligible_frontier_interval _ this]
-  show ?thesis
-    apply (rule *[of c d])
-    using box_subset_cbox[of c d]
-    apply auto
-    done
-qed
-
-lemma has_integral_restrict_closed_subintervals_eq:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach"
-  assumes "cbox c d \<subseteq> cbox a b"
-  shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b) \<longleftrightarrow> (f has_integral i) (cbox c d)"
-  (is "?l = ?r")
-proof (cases "cbox c d = {}")
-  case False
-  let ?g = "\<lambda>x. if x \<in> cbox c d then f x else 0"
-  show ?thesis
-    apply rule
-    defer
-    apply (rule has_integral_restrict_closed_subinterval[OF _ assms])
-    apply assumption
-  proof -
-    assume ?l
-    then have "?g integrable_on cbox c d"
-      using assms has_integral_integrable integrable_subinterval by blast
-    then have *: "f integrable_on cbox c d"
-      apply -
-      apply (rule integrable_eq)
-      apply auto
-      done
-    then have "i = integral (cbox c d) f"
-      apply -
-      apply (rule has_integral_unique)
-      apply (rule \<open>?l\<close>)
-      apply (rule has_integral_restrict_closed_subinterval[OF _ assms])
-      apply auto
-      done
-    then show ?r
-      using * by auto
-  qed
-qed auto
-
-
-text \<open>Hence we can apply the limit process uniformly to all integrals.\<close>
-
-lemma has_integral':
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "(f has_integral i) s \<longleftrightarrow>
-    (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      (\<exists>z. ((\<lambda>x. if x \<in> s then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - i) < e))"
-  (is "?l \<longleftrightarrow> (\<forall>e>0. ?r e)")
-proof -
-  {
-    presume *: "\<exists>a b. s = cbox a b \<Longrightarrow> ?thesis"
-    show ?thesis
-      apply cases
-      apply (rule *)
-      apply assumption
-      apply (subst has_integral_alt)
-      apply auto
-      done
-  }
-  assume "\<exists>a b. s = cbox a b"
-  then guess a b by (elim exE) note s=this
-  from bounded_cbox[of a b, unfolded bounded_pos] guess B ..
-  note B = conjunctD2[OF this,rule_format] show ?thesis
-    apply safe
-  proof -
-    fix e :: real
-    assume ?l and "e > 0"
-    show "?r e"
-      apply (rule_tac x="B+1" in exI)
-      apply safe
-      defer
-      apply (rule_tac x=i in exI)
-    proof
-      fix c d :: 'n
-      assume as: "ball 0 (B+1) \<subseteq> cbox c d"
-      then show "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) (cbox c d)"
-        unfolding s
-        apply -
-        apply (rule has_integral_restrict_closed_subinterval)
-        apply (rule \<open>?l\<close>[unfolded s])
-        apply safe
-        apply (drule B(2)[rule_format])
-        unfolding subset_eq
-        apply (erule_tac x=x in ballE)
-        apply (auto simp add: dist_norm)
-        done
-    qed (insert B \<open>e>0\<close>, auto)
-  next
-    assume as: "\<forall>e>0. ?r e"
-    from this[rule_format,OF zero_less_one] guess C .. note C=conjunctD2[OF this,rule_format]
-    define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
-    define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
-    have c_d: "cbox a b \<subseteq> cbox c d"
-      apply safe
-      apply (drule B(2))
-      unfolding mem_box
-    proof
-      fix x i
-      show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm x \<le> B" and "i \<in> Basis"
-        using that and Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x]
-        unfolding c_def d_def
-        by (auto simp add: field_simps setsum_negf)
-    qed
-    have "ball 0 C \<subseteq> cbox c d"
-      apply (rule subsetI)
-      unfolding mem_box mem_ball dist_norm
-    proof
-      fix x i :: 'n
-      assume x: "norm (0 - x) < C" and i: "i \<in> Basis"
-      show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
-        using Basis_le_norm[OF i, of x] and x i
-        unfolding c_def d_def
-        by (auto simp: setsum_negf)
-    qed
-    from C(2)[OF this] have "\<exists>y. (f has_integral y) (cbox a b)"
-      unfolding has_integral_restrict_closed_subintervals_eq[OF c_d,symmetric]
-      unfolding s
-      by auto
-    then guess y .. note y=this
-
-    have "y = i"
-    proof (rule ccontr)
-      assume "\<not> ?thesis"
-      then have "0 < norm (y - i)"
-        by auto
-      from as[rule_format,OF this] guess C ..  note C=conjunctD2[OF this,rule_format]
-      define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)"
-      define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)"
-      have c_d: "cbox a b \<subseteq> cbox c d"
-        apply safe
-        apply (drule B(2))
-        unfolding mem_box
-      proof
-        fix x i :: 'n
-        assume "norm x \<le> B" and "i \<in> Basis"
-        then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
-          using Basis_le_norm[of i x]
-          unfolding c_def d_def
-          by (auto simp add: field_simps setsum_negf)
-      qed
-      have "ball 0 C \<subseteq> cbox c d"
-        apply (rule subsetI)
-        unfolding mem_box mem_ball dist_norm
-      proof
-        fix x i :: 'n
-        assume "norm (0 - x) < C" and "i \<in> Basis"
-        then show "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i"
-          using Basis_le_norm[of i x]
-          unfolding c_def d_def
-          by (auto simp: setsum_negf)
-      qed
-      note C(2)[OF this] then guess z .. note z = conjunctD2[OF this, unfolded s]
-      note this[unfolded has_integral_restrict_closed_subintervals_eq[OF c_d]]
-      then have "z = y" and "norm (z - i) < norm (y - i)"
-        apply -
-        apply (rule has_integral_unique[OF _ y(1)])
-        apply assumption
-        apply assumption
-        done
-      then show False
-        by auto
-    qed
-    then show ?l
-      using y
-      unfolding s
-      by auto
-  qed
-qed
-
-lemma has_integral_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "(f has_integral i) s"
-    and "(g has_integral j) s"
-    and "\<forall>x\<in>s. f x \<le> g x"
-  shows "i \<le> j"
-  using has_integral_component_le[OF _ assms(1-2), of 1]
-  using assms(3)
-  by auto
-
-lemma integral_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "f integrable_on s"
-    and "g integrable_on s"
-    and "\<forall>x\<in>s. f x \<le> g x"
-  shows "integral s f \<le> integral s g"
-  by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)])
-
-lemma has_integral_nonneg:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "(f has_integral i) s"
-    and "\<forall>x\<in>s. 0 \<le> f x"
-  shows "0 \<le> i"
-  using has_integral_component_nonneg[of 1 f i s]
-  unfolding o_def
-  using assms
-  by auto
-
-lemma integral_nonneg:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "f integrable_on s"
-    and "\<forall>x\<in>s. 0 \<le> f x"
-  shows "0 \<le> integral s f"
-  by (rule has_integral_nonneg[OF assms(1)[unfolded has_integral_integral] assms(2)])
-
-
-text \<open>Hence a general restriction property.\<close>
-
-lemma has_integral_restrict[simp]:
-  assumes "s \<subseteq> t"
-  shows "((\<lambda>x. if x \<in> s then f x else (0::'a::banach)) has_integral i) t \<longleftrightarrow> (f has_integral i) s"
-proof -
-  have *: "\<And>x. (if x \<in> t then if x \<in> s then f x else 0 else 0) =  (if x\<in>s then f x else 0)"
-    using assms by auto
-  show ?thesis
-    apply (subst(2) has_integral')
-    apply (subst has_integral')
-    unfolding *
-    apply rule
-    done
-qed
-
-lemma has_integral_restrict_univ:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) UNIV \<longleftrightarrow> (f has_integral i) s"
-  by auto
-
-lemma has_integral_on_superset:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
-    and "s \<subseteq> t"
-    and "(f has_integral i) s"
-  shows "(f has_integral i) t"
-proof -
-  have "(\<lambda>x. if x \<in> s then f x else 0) = (\<lambda>x. if x \<in> t then f x else 0)"
-    apply rule
-    using assms(1-2)
-    apply auto
-    done
-  then show ?thesis
-    using assms(3)
-    apply (subst has_integral_restrict_univ[symmetric])
-    apply (subst(asm) has_integral_restrict_univ[symmetric])
-    apply auto
-    done
-qed
-
-lemma integrable_on_superset:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "\<forall>x. x \<notin> s \<longrightarrow> f x = 0"
-    and "s \<subseteq> t"
-    and "f integrable_on s"
-  shows "f integrable_on t"
-  using assms
-  unfolding integrable_on_def
-  by (auto intro:has_integral_on_superset)
-
-lemma integral_restrict_univ[intro]:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "f integrable_on s \<Longrightarrow> integral UNIV (\<lambda>x. if x \<in> s then f x else 0) = integral s f"
-  apply (rule integral_unique)
-  unfolding has_integral_restrict_univ
-  apply auto
-  done
-
-lemma integrable_restrict_univ:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "(\<lambda>x. if x \<in> s then f x else 0) integrable_on UNIV \<longleftrightarrow> f integrable_on s"
-  unfolding integrable_on_def
-  by auto
-
-lemma negligible_on_intervals: "negligible s \<longleftrightarrow> (\<forall>a b. negligible(s \<inter> cbox a b))" (is "?l \<longleftrightarrow> ?r")
-proof
-  assume ?r
-  show ?l
-    unfolding negligible_def
-  proof safe
-    fix a b
-    show "(indicator s has_integral 0) (cbox a b)"
-      apply (rule has_integral_negligible[OF \<open>?r\<close>[rule_format,of a b]])
-      unfolding indicator_def
-      apply auto
-      done
-  qed
-qed auto
-
-lemma negligible_translation:
-  assumes "negligible S"
-    shows "negligible (op + c ` S)"
-proof -
-  have inj: "inj (op + c)"
-    by simp
-  show ?thesis
-  using assms
-  proof (clarsimp simp: negligible_def)
-    fix a b
-    assume "\<forall>x y. (indicator S has_integral 0) (cbox x y)"
-    then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))"
-      by (meson Diff_iff assms has_integral_negligible indicator_simps(2))
-    have eq: "indicator (op + c ` S) = (\<lambda>x. indicator S (x - c))"
-      by (force simp add: indicator_def)
-    show "(indicator (op + c ` S) has_integral 0) (cbox a b)"
-      using has_integral_affinity [OF *, of 1 "-c"]
-            cbox_translation [of "c" "-c+a" "-c+b"]
-      by (simp add: eq add.commute)
-  qed
-qed
-
-lemma negligible_translation_rev:
-  assumes "negligible (op + c ` S)"
-    shows "negligible S"
-by (metis negligible_translation [OF assms, of "-c"] translation_galois)
-
-lemma has_integral_spike_set_eq:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "negligible ((s - t) \<union> (t - s))"
-  shows "(f has_integral y) s \<longleftrightarrow> (f has_integral y) t"
-  unfolding has_integral_restrict_univ[symmetric,of f]
-  apply (rule has_integral_spike_eq[OF assms])
-  by (auto split: if_split_asm)
-
-lemma has_integral_spike_set[dest]:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "negligible ((s - t) \<union> (t - s))"
-    and "(f has_integral y) s"
-  shows "(f has_integral y) t"
-  using assms has_integral_spike_set_eq
-  by auto
-
-lemma integrable_spike_set[dest]:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "negligible ((s - t) \<union> (t - s))"
-    and "f integrable_on s"
-  shows "f integrable_on t"
-  using assms(2)
-  unfolding integrable_on_def
-  unfolding has_integral_spike_set_eq[OF assms(1)] .
-
-lemma integrable_spike_set_eq:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "negligible ((s - t) \<union> (t - s))"
-  shows "f integrable_on s \<longleftrightarrow> f integrable_on t"
-  apply rule
-  apply (rule_tac[!] integrable_spike_set)
-  using assms
-  apply auto
-  done
-
-(*lemma integral_spike_set:
- "\<forall>f:real^M->real^N g s t.
-        negligible(s DIFF t \<union> t DIFF s)
-        \<longrightarrow> integral s f = integral t f"
-qed  REPEAT STRIP_TAC THEN REWRITE_TAC[integral] THEN
-  AP_TERM_TAC THEN ABS_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  ASM_MESON_TAC[]);;
-
-lemma has_integral_interior:
- "\<forall>f:real^M->real^N y s.
-        negligible(frontier s)
-        \<longrightarrow> ((f has_integral y) (interior s) \<longleftrightarrow> (f has_integral y) s)"
-qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
-    NEGLIGIBLE_SUBSET)) THEN
-  REWRITE_TAC[frontier] THEN
-  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
-  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
-  SET_TAC[]);;
-
-lemma has_integral_closure:
- "\<forall>f:real^M->real^N y s.
-        negligible(frontier s)
-        \<longrightarrow> ((f has_integral y) (closure s) \<longleftrightarrow> (f has_integral y) s)"
-qed  REPEAT STRIP_TAC THEN MATCH_MP_TAC HAS_INTEGRAL_SPIKE_SET_EQ THEN
-  FIRST_X_ASSUM(MATCH_MP_TAC o MATCH_MP (REWRITE_RULE[IMP_CONJ]
-    NEGLIGIBLE_SUBSET)) THEN
-  REWRITE_TAC[frontier] THEN
-  MP_TAC(ISPEC `s:real^M->bool` INTERIOR_SUBSET) THEN
-  MP_TAC(ISPEC `s:real^M->bool` CLOSURE_SUBSET) THEN
-  SET_TAC[]);;*)
-
-
-subsection \<open>More lemmas that are useful later\<close>
-
-lemma has_integral_subset_component_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes k: "k \<in> Basis"
-    and as: "s \<subseteq> t" "(f has_integral i) s" "(f has_integral j) t" "\<forall>x\<in>t. 0 \<le> f(x)\<bullet>k"
-  shows "i\<bullet>k \<le> j\<bullet>k"
-proof -
-  note has_integral_restrict_univ[symmetric, of f]
-  note as(2-3)[unfolded this] note * = has_integral_component_le[OF k this]
-  show ?thesis
-    apply (rule *)
-    using as(1,4)
-    apply auto
-    done
-qed
-
-lemma has_integral_subset_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "s \<subseteq> t"
-    and "(f has_integral i) s"
-    and "(f has_integral j) t"
-    and "\<forall>x\<in>t. 0 \<le> f x"
-  shows "i \<le> j"
-  using has_integral_subset_component_le[OF _ assms(1), of 1 f i j]
-  using assms
-  by auto
-
-lemma integral_subset_component_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "k \<in> Basis"
-    and "s \<subseteq> t"
-    and "f integrable_on s"
-    and "f integrable_on t"
-    and "\<forall>x \<in> t. 0 \<le> f x \<bullet> k"
-  shows "(integral s f)\<bullet>k \<le> (integral t f)\<bullet>k"
-  apply (rule has_integral_subset_component_le)
-  using assms
-  apply auto
-  done
-
-lemma integral_subset_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "s \<subseteq> t"
-    and "f integrable_on s"
-    and "f integrable_on t"
-    and "\<forall>x \<in> t. 0 \<le> f x"
-  shows "integral s f \<le> integral t f"
-  apply (rule has_integral_subset_le)
-  using assms
-  apply auto
-  done
-
-lemma has_integral_alt':
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "(f has_integral i) s \<longleftrightarrow> (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
-    (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e)"
-  (is "?l = ?r")
-proof
-  assume ?r
-  show ?l
-    apply (subst has_integral')
-    apply safe
-  proof goal_cases
-    case (1 e)
-    from \<open>?r\<close>[THEN conjunct2,rule_format,OF this] guess B .. note B=conjunctD2[OF this]
-    show ?case
-      apply rule
-      apply rule
-      apply (rule B)
-      apply safe
-      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0)" in exI)
-      apply (drule B(2)[rule_format])
-      using integrable_integral[OF \<open>?r\<close>[THEN conjunct1,rule_format]]
-      apply auto
-      done
-  qed
-next
-  assume ?l note as = this[unfolded has_integral'[of f],rule_format]
-  let ?f = "\<lambda>x. if x \<in> s then f x else 0"
-  show ?r
-  proof safe
-    fix a b :: 'n
-    from as[OF zero_less_one] guess B .. note B=conjunctD2[OF this,rule_format]
-    let ?a = "\<Sum>i\<in>Basis. min (a\<bullet>i) (-B) *\<^sub>R i::'n"
-    let ?b = "\<Sum>i\<in>Basis. max (b\<bullet>i) B *\<^sub>R i::'n"
-    show "?f integrable_on cbox a b"
-    proof (rule integrable_subinterval[of _ ?a ?b])
-      have "ball 0 B \<subseteq> cbox ?a ?b"
-        apply (rule subsetI)
-        unfolding mem_ball mem_box dist_norm
-      proof (rule, goal_cases)
-        case (1 x i)
-        then show ?case using Basis_le_norm[of i x]
-          by (auto simp add:field_simps)
-      qed
-      from B(2)[OF this] guess z .. note conjunct1[OF this]
-      then show "?f integrable_on cbox ?a ?b"
-        unfolding integrable_on_def by auto
-      show "cbox a b \<subseteq> cbox ?a ?b"
-        apply safe
-        unfolding mem_box
-        apply rule
-        apply (erule_tac x=i in ballE)
-        apply auto
-        done
-    qed
-
-    fix e :: real
-    assume "e > 0"
-    from as[OF this] guess B .. note B=conjunctD2[OF this,rule_format]
-    show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow>
-      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
-      apply rule
-      apply rule
-      apply (rule B)
-      apply safe
-    proof goal_cases
-      case 1
-      from B(2)[OF this] guess z .. note z=conjunctD2[OF this]
-      from integral_unique[OF this(1)] show ?case
-        using z(2) by auto
-    qed
-  qed
-qed
-
-
-subsection \<open>Continuity of the integral (for a 1-dimensional interval).\<close>
-
-lemma integrable_alt:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "f integrable_on s \<longleftrightarrow>
-    (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and>
-    (\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
-    norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) -
-      integral (cbox c d)  (\<lambda>x. if x \<in> s then f x else 0)) < e)"
-  (is "?l = ?r")
-proof
-  assume ?l
-  then guess y unfolding integrable_on_def .. note this[unfolded has_integral_alt'[of f]]
-  note y=conjunctD2[OF this,rule_format]
-  show ?r
-    apply safe
-    apply (rule y)
-  proof goal_cases
-    case (1 e)
-    then have "e/2 > 0"
-      by auto
-    from y(2)[OF this] guess B .. note B=conjunctD2[OF this,rule_format]
-    show ?case
-      apply rule
-      apply rule
-      apply (rule B)
-      apply safe
-    proof goal_cases
-      case prems: (1 a b c d)
-      show ?case
-        apply (rule norm_triangle_half_l)
-        using B(2)[OF prems(1)] B(2)[OF prems(2)]
-        apply auto
-        done
-    qed
-  qed
-next
-  assume ?r
-  note as = conjunctD2[OF this,rule_format]
-  let ?cube = "\<lambda>n. cbox (\<Sum>i\<in>Basis. - real n *\<^sub>R i::'n) (\<Sum>i\<in>Basis. real n *\<^sub>R i)"
-  have "Cauchy (\<lambda>n. integral (?cube n) (\<lambda>x. if x \<in> s then f x else 0))"
-  proof (unfold Cauchy_def, safe, goal_cases)
-    case (1 e)
-    from as(2)[OF this] guess B .. note B = conjunctD2[OF this,rule_format]
-    from real_arch_simple[of B] guess N .. note N = this
-    {
-      fix n
-      assume n: "n \<ge> N"
-      have "ball 0 B \<subseteq> ?cube n"
-        apply (rule subsetI)
-        unfolding mem_ball mem_box dist_norm
-      proof (rule, goal_cases)
-        case (1 x i)
-        then show ?case
-          using Basis_le_norm[of i x] \<open>i\<in>Basis\<close>
-          using n N
-          by (auto simp add: field_simps setsum_negf)
-      qed
-    }
-    then show ?case
-      apply -
-      apply (rule_tac x=N in exI)
-      apply safe
-      unfolding dist_norm
-      apply (rule B(2))
-      apply auto
-      done
-  qed
-  from this[unfolded convergent_eq_cauchy[symmetric]] guess i ..
-  note i = this[THEN LIMSEQ_D]
-
-  show ?l unfolding integrable_on_def has_integral_alt'[of f]
-    apply (rule_tac x=i in exI)
-    apply safe
-    apply (rule as(1)[unfolded integrable_on_def])
-  proof goal_cases
-    case (1 e)
-    then have *: "e/2 > 0" by auto
-    from i[OF this] guess N .. note N =this[rule_format]
-    from as(2)[OF *] guess B .. note B=conjunctD2[OF this,rule_format]
-    let ?B = "max (real N) B"
-    show ?case
-      apply (rule_tac x="?B" in exI)
-    proof safe
-      show "0 < ?B"
-        using B(1) by auto
-      fix a b :: 'n
-      assume ab: "ball 0 ?B \<subseteq> cbox a b"
-      from real_arch_simple[of ?B] guess n .. note n=this
-      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e"
-        apply (rule norm_triangle_half_l)
-        apply (rule B(2))
-        defer
-        apply (subst norm_minus_commute)
-        apply (rule N[of n])
-      proof safe
-        show "N \<le> n"
-          using n by auto
-        fix x :: 'n
-        assume x: "x \<in> ball 0 B"
-        then have "x \<in> ball 0 ?B"
-          by auto
-        then show "x \<in> cbox a b"
-          using ab by blast
-        show "x \<in> ?cube n"
-          using x
-          unfolding mem_box mem_ball dist_norm
-          apply -
-        proof (rule, goal_cases)
-          case (1 i)
-          then show ?case
-            using Basis_le_norm[of i x] \<open>i \<in> Basis\<close>
-            using n
-            by (auto simp add: field_simps setsum_negf)
-        qed
-      qed
-    qed
-  qed
-qed
-
-lemma integrable_altD:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s"
-  shows "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b"
-    and "\<And>e. e > 0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow>
-      norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d)  (\<lambda>x. if x \<in> s then f x else 0)) < e"
-  using assms[unfolded integrable_alt[of f]] by auto
-
-lemma integrable_on_subcbox:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s"
-    and "cbox a b \<subseteq> s"
-  shows "f integrable_on cbox a b"
-  apply (rule integrable_eq)
-  defer
-  apply (rule integrable_altD(1)[OF assms(1)])
-  using assms(2)
-  apply auto
-  done
-
-
-subsection \<open>A straddling criterion for integrability\<close>
-
-lemma integrable_straddle_interval:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "\<forall>e>0. \<exists>g  h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and>
-    norm (i - j) < e \<and> (\<forall>x\<in>cbox a b. (g x) \<le> f x \<and> f x \<le> h x)"
-  shows "f integrable_on cbox a b"
-proof (subst integrable_cauchy, safe, goal_cases)
-  case (1 e)
-  then have e: "e/3 > 0"
-    by auto
-  note assms[rule_format,OF this]
-  then guess g h i j by (elim exE conjE) note obt = this
-  from obt(1)[unfolded has_integral[of g], rule_format, OF e] guess d1 .. note d1=conjunctD2[OF this,rule_format]
-  from obt(2)[unfolded has_integral[of h], rule_format, OF e] guess d2 .. note d2=conjunctD2[OF this,rule_format]
-  show ?case
-    apply (rule_tac x="\<lambda>x. d1 x \<inter> d2 x" in exI)
-    apply (rule conjI gauge_inter d1 d2)+
-    unfolding fine_inter
-  proof (safe, goal_cases)
-    have **: "\<And>i j g1 g2 h1 h2 f1 f2. g1 - h2 \<le> f1 - f2 \<Longrightarrow> f1 - f2 \<le> h1 - g2 \<Longrightarrow>
-      \<bar>i - j\<bar> < e / 3 \<Longrightarrow> \<bar>g2 - i\<bar> < e / 3 \<Longrightarrow> \<bar>g1 - i\<bar> < e / 3 \<Longrightarrow>
-      \<bar>h2 - j\<bar> < e / 3 \<Longrightarrow> \<bar>h1 - j\<bar> < e / 3 \<Longrightarrow> \<bar>f1 - f2\<bar> < e"
-    using \<open>e > 0\<close> by arith
-    case prems: (1 p1 p2)
-    note tagged_division_ofD(2-4) note * = this[OF prems(1)] this[OF prems(4)]
-
-    have "(\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R g x) \<ge> 0"
-      and "0 \<le> (\<Sum>(x, k)\<in>p2. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)"
-      and "(\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R g x) \<ge> 0"
-      and "0 \<le> (\<Sum>(x, k)\<in>p1. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x)"
-      unfolding setsum_subtractf[symmetric]
-      apply -
-      apply (rule_tac[!] setsum_nonneg)
-      apply safe
-      unfolding real_scaleR_def right_diff_distrib[symmetric]
-      apply (rule_tac[!] mult_nonneg_nonneg)
-    proof -
-      fix a b
-      assume ab: "(a, b) \<in> p1"
-      show "0 \<le> content b"
-        using *(3)[OF ab]
-        apply safe
-        apply (rule content_pos_le)
-        done
-      then show "0 \<le> content b" .
-      show "0 \<le> f a - g a" "0 \<le> h a - f a"
-        using *(1-2)[OF ab]
-        using obt(4)[rule_format,of a]
-        by auto
-    next
-      fix a b
-      assume ab: "(a, b) \<in> p2"
-      show "0 \<le> content b"
-        using *(6)[OF ab]
-        apply safe
-        apply (rule content_pos_le)
-        done
-      then show "0 \<le> content b" .
-      show "0 \<le> f a - g a" and "0 \<le> h a - f a"
-        using *(4-5)[OF ab] using obt(4)[rule_format,of a] by auto
-    qed
-    then show ?case
-      apply -
-      unfolding real_norm_def
-      apply (rule **)
-      defer
-      defer
-      unfolding real_norm_def[symmetric]
-      apply (rule obt(3))
-      apply (rule d1(2)[OF conjI[OF prems(4,5)]])
-      apply (rule d1(2)[OF conjI[OF prems(1,2)]])
-      apply (rule d2(2)[OF conjI[OF prems(4,6)]])
-      apply (rule d2(2)[OF conjI[OF prems(1,3)]])
-      apply auto
-      done
-  qed
-qed
-
-lemma integrable_straddle:
-  fixes f :: "'n::euclidean_space \<Rightarrow> real"
-  assumes "\<forall>e>0. \<exists>g h i j. (g has_integral i) s \<and> (h has_integral j) s \<and>
-    norm (i - j) < e \<and> (\<forall>x\<in>s. g x \<le> f x \<and> f x \<le> h x)"
-  shows "f integrable_on s"
-proof -
-  have "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b"
-  proof (rule integrable_straddle_interval, safe, goal_cases)
-    case (1 a b e)
-    then have *: "e/4 > 0"
-      by auto
-    from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this
-    note obt(1)[unfolded has_integral_alt'[of g]]
-    note conjunctD2[OF this, rule_format]
-    note g = this(1) and this(2)[OF *]
-    from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
-    note obt(2)[unfolded has_integral_alt'[of h]]
-    note conjunctD2[OF this, rule_format]
-    note h = this(1) and this(2)[OF *]
-    from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
-    define c :: 'n where "c = (\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max B1 B2)) *\<^sub>R i)"
-    define d :: 'n where "d = (\<Sum>i\<in>Basis. max (b\<bullet>i) (max B1 B2) *\<^sub>R i)"
-    have *: "ball 0 B1 \<subseteq> cbox c d" "ball 0 B2 \<subseteq> cbox c d"
-      apply safe
-      unfolding mem_ball mem_box dist_norm
-      apply (rule_tac[!] ballI)
-    proof goal_cases
-      case (1 x i)
-      then show ?case using Basis_le_norm[of i x]
-        unfolding c_def d_def by auto
-    next
-      case (2 x i)
-      then show ?case using Basis_le_norm[of i x]
-        unfolding c_def d_def by auto
-    qed
-    have **: "\<And>ch cg ag ah::real. norm (ah - ag) \<le> norm (ch - cg) \<Longrightarrow> norm (cg - i) < e / 4 \<Longrightarrow>
-      norm (ch - j) < e / 4 \<Longrightarrow> norm (ag - ah) < e"
-      using obt(3)
-      unfolding real_norm_def
-      by arith
-    show ?case
-      apply (rule_tac x="\<lambda>x. if x \<in> s then g x else 0" in exI)
-      apply (rule_tac x="\<lambda>x. if x \<in> s then h x else 0" in exI)
-      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)" in exI)
-      apply (rule_tac x="integral (cbox a b) (\<lambda>x. if x \<in> s then h x else 0)" in exI)
-      apply safe
-      apply (rule_tac[1-2] integrable_integral,rule g)
-      apply (rule h)
-      apply (rule **[OF _ B1(2)[OF *(1)] B2(2)[OF *(2)]])
-    proof -
-      have *: "\<And>x f g. (if x \<in> s then f x else 0) - (if x \<in> s then g x else 0) =
-        (if x \<in> s then f x - g x else (0::real))"
-        by auto
-      note ** = abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF h g]]
-      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then h x else 0) -
-          integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)) \<le>
-        norm (integral (cbox c d) (\<lambda>x. if x \<in> s then h x else 0) -
-          integral (cbox c d) (\<lambda>x. if x \<in> s then g x else 0))"
-        unfolding integral_diff[OF h g,symmetric] real_norm_def
-        apply (subst **)
-        defer
-        apply (subst **)
-        defer
-        apply (rule has_integral_subset_le)
-        defer
-        apply (rule integrable_integral integrable_diff h g)+
-      proof safe
-        fix x
-        assume "x \<in> cbox a b"
-        then show "x \<in> cbox c d"
-          unfolding mem_box c_def d_def
-          apply -
-          apply rule
-          apply (erule_tac x=i in ballE)
-          apply auto
-          done
-      qed (insert obt(4), auto)
-    qed (insert obt(4), auto)
-  qed
-  note interv = this
-
-  show ?thesis
-    unfolding integrable_alt[of f]
-    apply safe
-    apply (rule interv)
-  proof goal_cases
-    case (1 e)
-    then have *: "e/3 > 0"
-      by auto
-    from assms[rule_format,OF this] guess g h i j by (elim exE conjE) note obt=this
-    note obt(1)[unfolded has_integral_alt'[of g]]
-    note conjunctD2[OF this, rule_format]
-    note g = this(1) and this(2)[OF *]
-    from this(2) guess B1 .. note B1 = conjunctD2[OF this,rule_format]
-    note obt(2)[unfolded has_integral_alt'[of h]]
-    note conjunctD2[OF this, rule_format]
-    note h = this(1) and this(2)[OF *]
-    from this(2) guess B2 .. note B2 = conjunctD2[OF this,rule_format]
-    show ?case
-      apply (rule_tac x="max B1 B2" in exI)
-      apply safe
-      apply (rule max.strict_coboundedI1)
-      apply (rule B1)
-    proof -
-      fix a b c d :: 'n
-      assume as: "ball 0 (max B1 B2) \<subseteq> cbox a b" "ball 0 (max B1 B2) \<subseteq> cbox c d"
-      have **: "ball 0 B1 \<subseteq> ball (0::'n) (max B1 B2)" "ball 0 B2 \<subseteq> ball (0::'n) (max B1 B2)"
-        by auto
-      have *: "\<And>ga gc ha hc fa fc::real.
-        \<bar>ga - i\<bar> < e / 3 \<and> \<bar>gc - i\<bar> < e / 3 \<and> \<bar>ha - j\<bar> < e / 3 \<and>
-        \<bar>hc - j\<bar> < e / 3 \<and> \<bar>i - j\<bar> < e / 3 \<and> ga \<le> fa \<and> fa \<le> ha \<and> gc \<le> fc \<and> fc \<le> hc \<Longrightarrow>
-        \<bar>fa - fc\<bar> < e"
-        by (simp add: abs_real_def split: if_split_asm)
-      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d)
-        (\<lambda>x. if x \<in> s then f x else 0)) < e"
-        unfolding real_norm_def
-        apply (rule *)
-        apply safe
-        unfolding real_norm_def[symmetric]
-        apply (rule B1(2))
-        apply (rule order_trans)
-        apply (rule **)
-        apply (rule as(1))
-        apply (rule B1(2))
-        apply (rule order_trans)
-        apply (rule **)
-        apply (rule as(2))
-        apply (rule B2(2))
-        apply (rule order_trans)
-        apply (rule **)
-        apply (rule as(1))
-        apply (rule B2(2))
-        apply (rule order_trans)
-        apply (rule **)
-        apply (rule as(2))
-        apply (rule obt)
-        apply (rule_tac[!] integral_le)
-        using obt
-        apply (auto intro!: h g interv)
-        done
-    qed
-  qed
-qed
-
-
-subsection \<open>Adding integrals over several sets\<close>
-
-lemma has_integral_union:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "(f has_integral i) s"
-    and "(f has_integral j) t"
-    and "negligible (s \<inter> t)"
-  shows "(f has_integral (i + j)) (s \<union> t)"
-proof -
-  note * = has_integral_restrict_univ[symmetric, of f]
-  show ?thesis
-    unfolding *
-    apply (rule has_integral_spike[OF assms(3)])
-    defer
-    apply (rule has_integral_add[OF assms(1-2)[unfolded *]])
-    apply auto
-    done
-qed
-
-lemma integrable_union:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
-  assumes "negligible (A \<inter> B)" "f integrable_on A" "f integrable_on B"
-  shows   "f integrable_on (A \<union> B)"
-proof -
-  from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B"
-     by (auto simp: integrable_on_def)
-  from has_integral_union[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def)
-qed
-
-lemma integrable_union':
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach"
-  assumes "f integrable_on A" "f integrable_on B" "negligible (A \<inter> B)" "C = A \<union> B"
-  shows   "f integrable_on C"
-  using integrable_union[of A B f] assms by simp
-
-lemma has_integral_unions:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "finite t"
-    and "\<forall>s\<in>t. (f has_integral (i s)) s"
-    and "\<forall>s\<in>t. \<forall>s'\<in>t. s \<noteq> s' \<longrightarrow> negligible (s \<inter> s')"
-  shows "(f has_integral (setsum i t)) (\<Union>t)"
-proof -
-  note * = has_integral_restrict_univ[symmetric, of f]
-  have **: "negligible (\<Union>((\<lambda>(a,b). a \<inter> b) ` {(a,b). a \<in> t \<and> b \<in> {y. y \<in> t \<and> a \<noteq> y}}))"
-    apply (rule negligible_Union)
-    apply (rule finite_imageI)
-    apply (rule finite_subset[of _ "t \<times> t"])
-    defer
-    apply (rule finite_cartesian_product[OF assms(1,1)])
-    using assms(3)
-    apply auto
-    done
-  note assms(2)[unfolded *]
-  note has_integral_setsum[OF assms(1) this]
-  then show ?thesis
-    unfolding *
-    apply -
-    apply (rule has_integral_spike[OF **])
-    defer
-    apply assumption
-    apply safe
-  proof goal_cases
-    case prems: (1 x)
-    then show ?case
-    proof (cases "x \<in> \<Union>t")
-      case True
-      then guess s unfolding Union_iff .. note s=this
-      then have *: "\<forall>b\<in>t. x \<in> b \<longleftrightarrow> b = s"
-        using prems(3) by blast
-      show ?thesis
-        unfolding if_P[OF True]
-        apply (rule trans)
-        defer
-        apply (rule setsum.cong)
-        apply (rule refl)
-        apply (subst *)
-        apply assumption
-        apply (rule refl)
-        unfolding setsum.delta[OF assms(1)]
-        using s
-        apply auto
-        done
-    qed auto
-  qed
-qed
-
-
-text \<open>In particular adding integrals over a division, maybe not of an interval.\<close>
-
-lemma has_integral_combine_division:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "d division_of s"
-    and "\<forall>k\<in>d. (f has_integral (i k)) k"
-  shows "(f has_integral (setsum i d)) s"
-proof -
-  note d = division_ofD[OF assms(1)]
-  show ?thesis
-    unfolding d(6)[symmetric]
-    apply (rule has_integral_unions)
-    apply (rule d assms)+
-    apply rule
-    apply rule
-    apply rule
-  proof goal_cases
-    case prems: (1 s s')
-    from d(4)[OF this(1)] d(4)[OF this(2)] guess a c b d by (elim exE) note obt=this
-    from d(5)[OF prems] show ?case
-      unfolding obt interior_cbox
-      apply -
-      apply (rule negligible_subset[of "(cbox a b-box a b) \<union> (cbox c d-box c d)"])
-      apply (rule negligible_Un negligible_frontier_interval)+
-      apply auto
-      done
-  qed
-qed
-
-lemma integral_combine_division_bottomup:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "d division_of s"
-    and "\<forall>k\<in>d. f integrable_on k"
-  shows "integral s f = setsum (\<lambda>i. integral i f) d"
-  apply (rule integral_unique)
-  apply (rule has_integral_combine_division[OF assms(1)])
-  using assms(2)
-  unfolding has_integral_integral
-  apply assumption
-  done
-
-lemma has_integral_combine_division_topdown:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s"
-    and "d division_of k"
-    and "k \<subseteq> s"
-  shows "(f has_integral (setsum (\<lambda>i. integral i f) d)) k"
-  apply (rule has_integral_combine_division[OF assms(2)])
-  apply safe
-  unfolding has_integral_integral[symmetric]
-proof goal_cases
-  case (1 k)
-  from division_ofD(2,4)[OF assms(2) this]
-  show ?case
-    apply safe
-    apply (rule integrable_on_subcbox)
-    apply (rule assms)
-    using assms(3)
-    apply auto
-    done
-qed
-
-lemma integral_combine_division_topdown:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s"
-    and "d division_of s"
-  shows "integral s f = setsum (\<lambda>i. integral i f) d"
-  apply (rule integral_unique)
-  apply (rule has_integral_combine_division_topdown)
-  using assms
-  apply auto
-  done
-
-lemma integrable_combine_division:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "d division_of s"
-    and "\<forall>i\<in>d. f integrable_on i"
-  shows "f integrable_on s"
-  using assms(2)
-  unfolding integrable_on_def
-  by (metis has_integral_combine_division[OF assms(1)])
-
-lemma integrable_on_subdivision:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "d division_of i"
-    and "f integrable_on s"
-    and "i \<subseteq> s"
-  shows "f integrable_on i"
-  apply (rule integrable_combine_division assms)+
-  apply safe
-proof goal_cases
-  case 1
-  note division_ofD(2,4)[OF assms(1) this]
-  then show ?case
-    apply safe
-    apply (rule integrable_on_subcbox[OF assms(2)])
-    using assms(3)
-    apply auto
-    done
-qed
-
-
-subsection \<open>Also tagged divisions\<close>
-
-lemma has_integral_combine_tagged_division:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "p tagged_division_of s"
-    and "\<forall>(x,k) \<in> p. (f has_integral (i k)) k"
-  shows "(f has_integral (setsum (\<lambda>(x,k). i k) p)) s"
-proof -
-  have *: "(f has_integral (setsum (\<lambda>k. integral k f) (snd ` p))) s"
-    apply (rule has_integral_combine_division)
-    apply (rule division_of_tagged_division[OF assms(1)])
-    using assms(2)
-    unfolding has_integral_integral[symmetric]
-    apply safe
-    apply auto
-    done
-  then show ?thesis
-    apply -
-    apply (rule subst[where P="\<lambda>i. (f has_integral i) s"])
-    defer
-    apply assumption
-    apply (rule trans[of _ "setsum (\<lambda>(x,k). integral k f) p"])
-    apply (subst eq_commute)
-    apply (rule setsum.over_tagged_division_lemma[OF assms(1)])
-    apply (rule integral_null)
-    apply assumption
-    apply (rule setsum.cong)
-    using assms(2)
-    apply auto
-    done
-qed
-
-lemma integral_combine_tagged_division_bottomup:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "p tagged_division_of (cbox a b)"
-    and "\<forall>(x,k)\<in>p. f integrable_on k"
-  shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
-  apply (rule integral_unique)
-  apply (rule has_integral_combine_tagged_division[OF assms(1)])
-  using assms(2)
-  apply auto
-  done
-
-lemma has_integral_combine_tagged_division_topdown:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on cbox a b"
-    and "p tagged_division_of (cbox a b)"
-  shows "(f has_integral (setsum (\<lambda>(x,k). integral k f) p)) (cbox a b)"
-  apply (rule has_integral_combine_tagged_division[OF assms(2)])
-  apply safe
-proof goal_cases
-  case 1
-  note tagged_division_ofD(3-4)[OF assms(2) this]
-  then show ?case
-    using integrable_subinterval[OF assms(1)] by blast
-qed
-
-lemma integral_combine_tagged_division_topdown:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on cbox a b"
-    and "p tagged_division_of (cbox a b)"
-  shows "integral (cbox a b) f = setsum (\<lambda>(x,k). integral k f) p"
-  apply (rule integral_unique)
-  apply (rule has_integral_combine_tagged_division_topdown)
-  using assms
-  apply auto
-  done
-
-
-subsection \<open>Henstock's lemma\<close>
-
-lemma henstock_lemma_part1:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on cbox a b"
-    and "e > 0"
-    and "gauge d"
-    and "(\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral(cbox a b) f) < e)"
-    and p: "p tagged_partial_division_of (cbox a b)" "d fine p"
-  shows "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x - integral k f) p) \<le> e"
-  (is "?x \<le> e")
-proof -
-  { presume "\<And>k. 0<k \<Longrightarrow> ?x \<le> e + k" then show ?thesis by (blast intro: field_le_epsilon) }
-  fix k :: real
-  assume k: "k > 0"
-  note p' = tagged_partial_division_ofD[OF p(1)]
-  have "\<Union>(snd ` p) \<subseteq> cbox a b"
-    using p'(3) by fastforce
-  note partial_division_of_tagged_division[OF p(1)] this
-  from partial_division_extend_interval[OF this] guess q . note q=this and q' = division_ofD[OF this(2)]
-  define r where "r = q - snd ` p"
-  have "snd ` p \<inter> r = {}"
-    unfolding r_def by auto
-  have r: "finite r"
-    using q' unfolding r_def by auto
-
-  have "\<forall>i\<in>r. \<exists>p. p tagged_division_of i \<and> d fine p \<and>
-    norm (setsum (\<lambda>(x,j). content j *\<^sub>R f x) p - integral i f) < k / (real (card r) + 1)"
-    apply safe
-  proof goal_cases
-    case (1 i)
-    then have i: "i \<in> q"
-      unfolding r_def by auto
-    from q'(4)[OF this] guess u v by (elim exE) note uv=this
-    have *: "k / (real (card r) + 1) > 0" using k by simp
-    have "f integrable_on cbox u v"
-      apply (rule integrable_subinterval[OF assms(1)])
-      using q'(2)[OF i]
-      unfolding uv
-      apply auto
-      done
-    note integrable_integral[OF this, unfolded has_integral[of f]]
-    from this[rule_format,OF *] guess dd .. note dd=conjunctD2[OF this,rule_format]
-    note gauge_inter[OF \<open>gauge d\<close> dd(1)]
-    from fine_division_exists[OF this,of u v] guess qq .
-    then show ?case
-      apply (rule_tac x=qq in exI)
-      using dd(2)[of qq]
-      unfolding fine_inter uv
-      apply auto
-      done
-  qed
-  from bchoice[OF this] guess qq .. note qq=this[rule_format]
-
-  let ?p = "p \<union> \<Union>(qq ` r)"
-  have "norm ((\<Sum>(x, k)\<in>?p. content k *\<^sub>R f x) - integral (cbox a b) f) < e"
-    apply (rule assms(4)[rule_format])
-  proof
-    show "d fine ?p"
-      apply (rule fine_union)
-      apply (rule p)
-      apply (rule fine_unions)
-      using qq
-      apply auto
-      done
-    note * = tagged_partial_division_of_union_self[OF p(1)]
-    have "p \<union> \<Union>(qq ` r) tagged_division_of \<Union>(snd ` p) \<union> \<Union>r"
-      using r
-    proof (rule tagged_division_union[OF * tagged_division_unions], goal_cases)
-      case 1
-      then show ?case
-        using qq by auto
-    next
-      case 2
-      then show ?case
-        apply rule
-        apply rule
-        apply rule
-        apply(rule q'(5))
-        unfolding r_def
-        apply auto
-        done
-    next
-      case 3
-      then show ?case
-        apply (rule inter_interior_unions_intervals)
-        apply fact
-        apply rule
-        apply rule
-        apply (rule q')
-        defer
-        apply rule
-        apply (subst Int_commute)
-        apply (rule inter_interior_unions_intervals)
-        apply (rule finite_imageI)
-        apply (rule p')
-        apply rule
-        defer
-        apply rule
-        apply (rule q')
-        using q(1) p'
-        unfolding r_def
-        apply auto
-        done
-    qed
-    moreover have "\<Union>(snd ` p) \<union> \<Union>r = cbox a b" and "{qq i |i. i \<in> r} = qq ` r"
-      unfolding Union_Un_distrib[symmetric] r_def
-      using q
-      by auto
-    ultimately show "?p tagged_division_of (cbox a b)"
-      by fastforce
-  qed
-
-  then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>\<Union>(qq ` r). content k *\<^sub>R f x) -
-    integral (cbox a b) f) < e"
-    apply (subst setsum.union_inter_neutral[symmetric])
-    apply (rule p')
-    prefer 3
-    apply assumption
-    apply rule
-    apply (rule r)
-    apply safe
-    apply (drule qq)
-  proof -
-    fix x l k
-    assume as: "(x, l) \<in> p" "(x, l) \<in> qq k" "k \<in> r"
-    note qq[OF this(3)]
-    note tagged_division_ofD(3,4)[OF conjunct1[OF this] as(2)]
-    from this(2) guess u v by (elim exE) note uv=this
-    have "l\<in>snd ` p" unfolding image_iff apply(rule_tac x="(x,l)" in bexI) using as by auto
-    then have "l \<in> q" "k \<in> q" "l \<noteq> k"
-      using as(1,3) q(1) unfolding r_def by auto
-    note q'(5)[OF this]
-    then have "interior l = {}"
-      using interior_mono[OF \<open>l \<subseteq> k\<close>] by blast
-    then show "content l *\<^sub>R f x = 0"
-      unfolding uv content_eq_0_interior[symmetric] by auto
-  qed auto
-
-  then have "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x))
-    (qq ` r) - integral (cbox a b) f) < e"
-    apply (subst (asm) setsum.Union_comp)
-    prefer 2
-    unfolding split_paired_all split_conv image_iff
-    apply (erule bexE)+
-  proof -
-    fix x m k l T1 T2
-    assume "(x, m) \<in> T1" "(x, m) \<in> T2" "T1 \<noteq> T2" "k \<in> r" "l \<in> r" "T1 = qq k" "T2 = qq l"
-    note as = this(1-5)[unfolded this(6-)]
-    note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]]
-    from this(2)[OF as(4,1)] guess u v by (elim exE) note uv=this
-    have *: "interior (k \<inter> l) = {}"
-      by (metis DiffE \<open>T1 = qq k\<close> \<open>T1 \<noteq> T2\<close> \<open>T2 = qq l\<close> as(4) as(5) interior_Int q'(5) r_def)
-    have "interior m = {}"
-      unfolding subset_empty[symmetric]
-      unfolding *[symmetric]
-      apply (rule interior_mono)
-      using kl(1)[OF as(4,1)] kl(1)[OF as(5,2)]
-      apply auto
-      done
-    then show "content m *\<^sub>R f x = 0"
-      unfolding uv content_eq_0_interior[symmetric]
-      by auto
-  qed (insert qq, auto)
-
-  then have **: "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) + setsum (setsum (\<lambda>(x, k). content k *\<^sub>R f x) \<circ> qq) r -
-    integral (cbox a b) f) < e"
-    apply (subst (asm) setsum.reindex_nontrivial)
-    apply fact
-    apply (rule setsum.neutral)
-    apply rule
-    unfolding split_paired_all split_conv
-    defer
-    apply assumption
-  proof -
-    fix k l x m
-    assume as: "k \<in> r" "l \<in> r" "k \<noteq> l" "qq k = qq l" "(x, m) \<in> qq k"
-    note tagged_division_ofD(6)[OF qq[THEN conjunct1]]
-    from this[OF as(1)] this[OF as(2)] show "content m *\<^sub>R f x = 0"
-      using as(3) unfolding as by auto
-  qed
-
-  have *: "norm (cp - ip) \<le> e + k"
-    if "norm ((cp + cr) - i) < e"
-    and "norm (cr - ir) < k"
-    and "ip + ir = i"
-    for ir ip i cr cp
-  proof -
-    from that show ?thesis
-      using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"]
-      unfolding that(3)[symmetric] norm_minus_cancel
-      by (auto simp add: algebra_simps)
-  qed
-
-  have "?x =  norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p. integral k f))"
-    unfolding split_def setsum_subtractf ..
-  also have "\<dots> \<le> e + k"
-    apply (rule *[OF **, where ir1="setsum (\<lambda>k. integral k f) r"])
-  proof goal_cases
-    case 1
-    have *: "k * real (card r) / (1 + real (card r)) < k"
-      using k by (auto simp add: field_simps)
-    show ?case
-      apply (rule le_less_trans[of _ "setsum (\<lambda>x. k / (real (card r) + 1)) r"])
-      unfolding setsum_subtractf[symmetric]
-      apply (rule setsum_norm_le)
-      apply rule
-      apply (drule qq)
-      defer
-      unfolding divide_inverse setsum_left_distrib[symmetric]
-      unfolding divide_inverse[symmetric]
-      using * apply (auto simp add: field_simps)
-      done
-  next
-    case 2
-    have *: "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)"
-      apply (subst setsum.reindex_nontrivial)
-      apply fact
-      unfolding split_paired_all snd_conv split_def o_def
-    proof -
-      fix x l y m
-      assume as: "(x, l) \<in> p" "(y, m) \<in> p" "(x, l) \<noteq> (y, m)" "l = m"
-      from p'(4)[OF as(1)] guess u v by (elim exE) note uv=this
-      show "integral l f = 0"
-        unfolding uv
-        apply (rule integral_unique)
-        apply (rule has_integral_null)
-        unfolding content_eq_0_interior
-        using p'(5)[OF as(1-3)]
-        unfolding uv as(4)[symmetric]
-        apply auto
-        done
-    qed auto
-    from q(1) have **: "snd ` p \<union> q = q" by auto
-    show ?case
-      unfolding integral_combine_division_topdown[OF assms(1) q(2)] * r_def
-      using ** q'(1) p'(1) setsum.union_disjoint [of "snd ` p" "q - snd ` p" "\<lambda>k. integral k f", symmetric]
-        by simp
-  qed
-  finally show "?x \<le> e + k" .
-qed
-
-lemma henstock_lemma_part2:
-  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
-  assumes "f integrable_on cbox a b"
-    and "e > 0"
-    and "gauge d"
-    and "\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-      norm (setsum (\<lambda>(x,k). content k *\<^sub>R f x) p - integral (cbox a b) f) < e"
-    and "p tagged_partial_division_of (cbox a b)"
-    and "d fine p"
-  shows "setsum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
-  unfolding split_def
-  apply (rule setsum_norm_allsubsets_bound)
-  defer
-  apply (rule henstock_lemma_part1[unfolded split_def,OF assms(1-3)])
-  apply safe
-  apply (rule assms[rule_format,unfolded split_def])
-  defer
-  apply (rule tagged_partial_division_subset)
-  apply (rule assms)
-  apply assumption
-  apply (rule fine_subset)
-  apply assumption
-  apply (rule assms)
-  using assms(5)
-  apply auto
-  done
-
-lemma henstock_lemma:
-  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
-  assumes "f integrable_on cbox a b"
-    and "e > 0"
-  obtains d where "gauge d"
-    and "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-      setsum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e"
-proof -
-  have *: "e / (2 * (real DIM('n) + 1)) > 0" using assms(2) by simp
-  from integrable_integral[OF assms(1),unfolded has_integral[of f],rule_format,OF this]
-  guess d .. note d = conjunctD2[OF this]
-  show thesis
-    apply (rule that)
-    apply (rule d)
-  proof (safe, goal_cases)
-    case (1 p)
-    note * = henstock_lemma_part2[OF assms(1) * d this]
-    show ?case
-      apply (rule le_less_trans[OF *])
-      using \<open>e > 0\<close>
-      apply (auto simp add: field_simps)
-      done
-  qed
-qed
-
-
-subsection \<open>Geometric progression\<close>
-
-text \<open>FIXME: Should one or more of these theorems be moved to @{file
-"~~/src/HOL/Set_Interval.thy"}, alongside \<open>geometric_sum\<close>?\<close>
-
-lemma sum_gp_basic:
-  fixes x :: "'a::ring_1"
-  shows "(1 - x) * setsum (\<lambda>i. x^i) {0 .. n} = (1 - x^(Suc n))"
-proof -
-  define y where "y = 1 - x"
-  have "y * (\<Sum>i=0..n. (1 - y) ^ i) = 1 - (1 - y) ^ Suc n"
-    by (induct n) (simp_all add: algebra_simps)
-  then show ?thesis
-    unfolding y_def by simp
-qed
-
-lemma sum_gp_multiplied:
-  assumes mn: "m \<le> n"
-  shows "((1::'a::{field}) - x) * setsum (op ^ x) {m..n} = x^m - x^ Suc n"
-  (is "?lhs = ?rhs")
-proof -
-  let ?S = "{0..(n - m)}"
-  from mn have mn': "n - m \<ge> 0"
-    by arith
-  let ?f = "op + m"
-  have i: "inj_on ?f ?S"
-    unfolding inj_on_def by auto
-  have f: "?f ` ?S = {m..n}"
-    using mn
-    apply (auto simp add: image_iff Bex_def)
-    apply presburger
-    done
-  have th: "op ^ x \<circ> op + m = (\<lambda>i. x^m * x^i)"
-    by (rule ext) (simp add: power_add power_mult)
-  from setsum.reindex[OF i, of "op ^ x", unfolded f th setsum_right_distrib[symmetric]]
-  have "?lhs = x^m * ((1 - x) * setsum (op ^ x) {0..n - m})"
-    by simp
-  then show ?thesis
-    unfolding sum_gp_basic
-    using mn
-    by (simp add: field_simps power_add[symmetric])
-qed
-
-lemma sum_gp:
-  "setsum (op ^ (x::'a::{field})) {m .. n} =
-    (if n < m then 0
-     else if x = 1 then of_nat ((n + 1) - m)
-     else (x^ m - x^ (Suc n)) / (1 - x))"
-proof -
-  {
-    assume nm: "n < m"
-    then have ?thesis by simp
-  }
-  moreover
-  {
-    assume "\<not> n < m"
-    then have nm: "m \<le> n"
-      by arith
-    {
-      assume x: "x = 1"
-      then have ?thesis
-        by simp
-    }
-    moreover
-    {
-      assume x: "x \<noteq> 1"
-      then have nz: "1 - x \<noteq> 0"
-        by simp
-      from sum_gp_multiplied[OF nm, of x] nz have ?thesis
-        by (simp add: field_simps)
-    }
-    ultimately have ?thesis by blast
-  }
-  ultimately show ?thesis by blast
-qed
-
-lemma sum_gp_offset:
-  "setsum (op ^ (x::'a::{field})) {m .. m+n} =
-    (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
-  unfolding sum_gp[of x m "m + n"] power_Suc
-  by (simp add: field_simps power_add)
-
-
-subsection \<open>Monotone convergence (bounded interval first)\<close>
-
-lemma monotone_convergence_interval:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
-  assumes "\<forall>k. (f k) integrable_on cbox a b"
-    and "\<forall>k. \<forall>x\<in>cbox a b.(f k x) \<le> f (Suc k) x"
-    and "\<forall>x\<in>cbox a b. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-    and "bounded {integral (cbox a b) (f k) | k . k \<in> UNIV}"
-  shows "g integrable_on cbox a b \<and> ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> integral (cbox a b) g) sequentially"
-proof (cases "content (cbox a b) = 0")
-  case True
-  show ?thesis
-    using integrable_on_null[OF True]
-    unfolding integral_null[OF True]
-    using tendsto_const
-    by auto
-next
-  case False
-  have fg: "\<forall>x\<in>cbox a b. \<forall>k. (f k x) \<bullet> 1 \<le> (g x) \<bullet> 1"
-  proof safe
-    fix x k
-    assume x: "x \<in> cbox a b"
-    note * = Lim_component_ge[OF assms(3)[rule_format, OF x] trivial_limit_sequentially]
-    show "f k x \<bullet> 1 \<le> g x \<bullet> 1"
-      apply (rule *)
-      unfolding eventually_sequentially
-      apply (rule_tac x=k in exI)
-      apply -
-      apply (rule transitive_stepwise_le)
-      using assms(2)[rule_format, OF x]
-      apply auto
-      done
-  qed
-  have "\<exists>i. ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> i) sequentially"
-    apply (rule bounded_increasing_convergent)
-    defer
-    apply rule
-    apply (rule integral_le)
-    apply safe
-    apply (rule assms(1-2)[rule_format])+
-    using assms(4)
-    apply auto
-    done
-  then guess i .. note i=this
-  have i': "\<And>k. (integral(cbox a b) (f k)) \<le> i\<bullet>1"
-    apply (rule Lim_component_ge)
-    apply (rule i)
-    apply (rule trivial_limit_sequentially)
-    unfolding eventually_sequentially
-    apply (rule_tac x=k in exI)
-    apply (rule transitive_stepwise_le)
-    prefer 3
-    unfolding inner_simps real_inner_1_right
-    apply (rule integral_le)
-    apply (rule assms(1-2)[rule_format])+
-    using assms(2)
-    apply auto
-    done
-
-  have "(g has_integral i) (cbox a b)"
-    unfolding has_integral
-  proof (safe, goal_cases)
-    case e: (1 e)
-    then have "\<forall>k. (\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow>
-      norm ((\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f k x) - integral (cbox a b) (f k)) < e / 2 ^ (k + 2)))"
-      apply -
-      apply rule
-      apply (rule assms(1)[unfolded has_integral_integral has_integral,rule_format])
-      apply auto
-      done
-    from choice[OF this] guess c .. note c=conjunctD2[OF this[rule_format],rule_format]
-
-    have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i\<bullet>1 - (integral (cbox a b) (f k)) \<and> i\<bullet>1 - (integral (cbox a b) (f k)) < e / 4"
-    proof -
-      have "e/4 > 0"
-        using e by auto
-      from LIMSEQ_D [OF i this] guess r ..
-      then show ?thesis
-        apply (rule_tac x=r in exI)
-        apply rule
-        apply (erule_tac x=k in allE)
-        subgoal for k using i'[of k] by auto
-        done
-    qed
-    then guess r .. note r=conjunctD2[OF this[rule_format]]
-
-    have "\<forall>x\<in>cbox a b. \<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x)\<bullet>1 - (f k x)\<bullet>1 \<and>
-      (g x)\<bullet>1 - (f k x)\<bullet>1 < e / (4 * content(cbox a b))"
-    proof (rule, goal_cases)
-      case prems: (1 x)
-      have "e / (4 * content (cbox a b)) > 0"
-        using \<open>e>0\<close> False content_pos_le[of a b] by auto
-      from assms(3)[rule_format, OF prems, THEN LIMSEQ_D, OF this]
-      guess n .. note n=this
-      then show ?case
-        apply (rule_tac x="n + r" in exI)
-        apply safe
-        apply (erule_tac[2-3] x=k in allE)
-        unfolding dist_real_def
-        using fg[rule_format, OF prems]
-        apply (auto simp add: field_simps)
-        done
-    qed
-    from bchoice[OF this] guess m .. note m=conjunctD2[OF this[rule_format],rule_format]
-    define d where "d x = c (m x) x" for x
-    show ?case
-      apply (rule_tac x=d in exI)
-    proof safe
-      show "gauge d"
-        using c(1) unfolding gauge_def d_def by auto
-    next
-      fix p
-      assume p: "p tagged_division_of (cbox a b)" "d fine p"
-      note p'=tagged_division_ofD[OF p(1)]
-      have "\<exists>a. \<forall>x\<in>p. m (fst x) \<le> a"
-        by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI)
-      then guess s .. note s=this
-      have *: "\<forall>a b c d. norm(a - b) \<le> e / 4 \<and> norm(b - c) < e / 2 \<and>
-        norm (c - d) < e / 4 \<longrightarrow> norm (a - d) < e"
-      proof (safe, goal_cases)
-        case (1 a b c d)
-        then show ?case
-          using norm_triangle_lt[of "a - b" "b - c" "3* e/4"]
-            norm_triangle_lt[of "a - b + (b - c)" "c - d" e]
-          unfolding norm_minus_cancel
-          by (auto simp add: algebra_simps)
-      qed
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - i) < e"
-        apply (rule *[rule_format,where
-          b="\<Sum>(x, k)\<in>p. content k *\<^sub>R f (m x) x" and c="\<Sum>(x, k)\<in>p. integral k (f (m x))"])
-      proof (safe, goal_cases)
-        case 1
-        show ?case
-          apply (rule order_trans[of _ "\<Sum>(x, k)\<in>p. content k * (e / (4 * content (cbox a b)))"])
-          unfolding setsum_subtractf[symmetric]
-          apply (rule order_trans)
-          apply (rule norm_setsum)
-          apply (rule setsum_mono)
-          unfolding split_paired_all split_conv
-          unfolding split_def setsum_left_distrib[symmetric] scaleR_diff_right[symmetric]
-          unfolding additive_content_tagged_division[OF p(1), unfolded split_def]
-        proof -
-          fix x k
-          assume xk: "(x, k) \<in> p"
-          then have x: "x \<in> cbox a b"
-            using p'(2-3)[OF xk] by auto
-          from p'(4)[OF xk] guess u v by (elim exE) note uv=this
-          show "norm (content k *\<^sub>R (g x - f (m x) x)) \<le> content k * (e / (4 * content (cbox a b)))"
-            unfolding norm_scaleR uv
-            unfolding abs_of_nonneg[OF content_pos_le]
-            apply (rule mult_left_mono)
-            using m(2)[OF x,of "m x"]
-            apply auto
-            done
-        qed (insert False, auto)
-
-      next
-        case 2
-        show ?case
-          apply (rule le_less_trans[of _ "norm (\<Sum>j = 0..s.
-            \<Sum>(x, k)\<in>{xk\<in>p. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))"])
-          apply (subst setsum_group)
-          apply fact
-          apply (rule finite_atLeastAtMost)
-          defer
-          apply (subst split_def)+
-          unfolding setsum_subtractf
-          apply rule
-        proof -
-          show "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> p.
-            m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e / 2"
-            apply (rule le_less_trans[of _ "setsum (\<lambda>i. e / 2^(i+2)) {0..s}"])
-            apply (rule setsum_norm_le)
-          proof
-            show "(\<Sum>i = 0..s. e / 2 ^ (i + 2)) < e / 2"
-              unfolding power_add divide_inverse inverse_mult_distrib
-              unfolding setsum_right_distrib[symmetric] setsum_left_distrib[symmetric]
-              unfolding power_inverse [symmetric] sum_gp
-              apply(rule mult_strict_left_mono[OF _ e])
-              unfolding power2_eq_square
-              apply auto
-              done
-            fix t
-            assume "t \<in> {0..s}"
-            show "norm (\<Sum>(x, k)\<in>{xk \<in> p. m (fst xk) = t}. content k *\<^sub>R f (m x) x -
-              integral k (f (m x))) \<le> e / 2 ^ (t + 2)"
-              apply (rule order_trans
-                [of _ "norm (setsum (\<lambda>(x,k). content k *\<^sub>R f t x - integral k (f t)) {xk \<in> p. m (fst xk) = t})"])
-              apply (rule eq_refl)
-              apply (rule arg_cong[where f=norm])
-              apply (rule setsum.cong)
-              apply (rule refl)
-              defer
-              apply (rule henstock_lemma_part1)
-              apply (rule assms(1)[rule_format])
-              apply (simp add: e)
-              apply safe
-              apply (rule c)+
-              apply rule
-              apply assumption+
-              apply (rule tagged_partial_division_subset[of p])
-              apply (rule p(1)[unfolded tagged_division_of_def,THEN conjunct1])
-              defer
-              unfolding fine_def
-              apply safe
-              apply (drule p(2)[unfolded fine_def,rule_format])
-              unfolding d_def
-              apply auto
-              done
-          qed
-        qed (insert s, auto)
-      next
-        case 3
-        note comb = integral_combine_tagged_division_topdown[OF assms(1)[rule_format] p(1)]
-        have *: "\<And>sr sx ss ks kr::real. kr = sr \<longrightarrow> ks = ss \<longrightarrow>
-          ks \<le> i \<and> sr \<le> sx \<and> sx \<le> ss \<and> 0 \<le> i\<bullet>1 - kr\<bullet>1 \<and> i\<bullet>1 - kr\<bullet>1 < e/4 \<longrightarrow> \<bar>sx - i\<bar> < e/4"
-          by auto
-        show ?case
-          unfolding real_norm_def
-          apply (rule *[rule_format])
-          apply safe
-          apply (rule comb[of r])
-          apply (rule comb[of s])
-          apply (rule i'[unfolded real_inner_1_right])
-          apply (rule_tac[1-2] setsum_mono)
-          unfolding split_paired_all split_conv
-          apply (rule_tac[1-2] integral_le[OF ])
-        proof safe
-          show "0 \<le> i\<bullet>1 - (integral (cbox a b) (f r))\<bullet>1"
-            using r(1) by auto
-          show "i\<bullet>1 - (integral (cbox a b) (f r))\<bullet>1 < e / 4"
-            using r(2) by auto
-          fix x k
-          assume xk: "(x, k) \<in> p"
-          from p'(4)[OF this] guess u v by (elim exE) note uv=this
-          show "f r integrable_on k"
-            and "f s integrable_on k"
-            and "f (m x) integrable_on k"
-            and "f (m x) integrable_on k"
-            unfolding uv
-            apply (rule_tac[!] integrable_on_subcbox[OF assms(1)[rule_format]])
-            using p'(3)[OF xk]
-            unfolding uv
-            apply auto
-            done
-          fix y
-          assume "y \<in> k"
-          then have "y \<in> cbox a b"
-            using p'(3)[OF xk] by auto
-          then have *: "\<And>m. \<forall>n\<ge>m. f m y \<le> f n y"
-            apply -
-            apply (rule transitive_stepwise_le)
-            using assms(2)
-            apply auto
-            done
-          show "f r y \<le> f (m x) y" and "f (m x) y \<le> f s y"
-            apply (rule_tac[!] *[rule_format])
-            using s[rule_format,OF xk] m(1)[of x] p'(2-3)[OF xk]
-            apply auto
-            done
-        qed
-      qed
-    qed
-  qed note * = this
-
-  have "integral (cbox a b) g = i"
-    by (rule integral_unique) (rule *)
-  then show ?thesis
-    using i * by auto
-qed
-
-lemma monotone_convergence_increasing:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
-  assumes "\<forall>k. (f k) integrable_on s"
-    and "\<forall>k. \<forall>x\<in>s. (f k x) \<le> (f (Suc k) x)"
-    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-    and "bounded {integral s (f k)| k. True}"
-  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-proof -
-  have lem: "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-    if "\<forall>k. \<forall>x\<in>s. 0 \<le> f k x"
-    and "\<forall>k. (f k) integrable_on s"
-    and "\<forall>k. \<forall>x\<in>s. f k x \<le> f (Suc k) x"
-    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-    and "bounded {integral s (f k)| k. True}"
-    for f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" and g s
-  proof -
-    note assms=that[rule_format]
-    have "\<forall>x\<in>s. \<forall>k. (f k x)\<bullet>1 \<le> (g x)\<bullet>1"
-      apply safe
-      apply (rule Lim_component_ge)
-      apply (rule that(4)[rule_format])
-      apply assumption
-      apply (rule trivial_limit_sequentially)
-      unfolding eventually_sequentially
-      apply (rule_tac x=k in exI)
-      apply (rule transitive_stepwise_le)
-      using that(3)
-      apply auto
-      done
-    note fg=this[rule_format]
-
-    have "\<exists>i. ((\<lambda>k. integral s (f k)) \<longlongrightarrow> i) sequentially"
-      apply (rule bounded_increasing_convergent)
-      apply (rule that(5))
-      apply rule
-      apply (rule integral_le)
-      apply (rule that(2)[rule_format])+
-      using that(3)
-      apply auto
-      done
-    then guess i .. note i=this
-    have "\<And>k. \<forall>x\<in>s. \<forall>n\<ge>k. f k x \<le> f n x"
-      apply rule
-      apply (rule transitive_stepwise_le)
-      using that(3)
-      apply auto
-      done
-    then have i': "\<forall>k. (integral s (f k))\<bullet>1 \<le> i\<bullet>1"
-      apply -
-      apply rule
-      apply (rule Lim_component_ge)
-      apply (rule i)
-      apply (rule trivial_limit_sequentially)
-      unfolding eventually_sequentially
-      apply (rule_tac x=k in exI)
-      apply safe
-      apply (rule integral_component_le)
-      apply simp
-      apply (rule that(2)[rule_format])+
-      apply auto
-      done
-
-    note int = assms(2)[unfolded integrable_alt[of _ s],THEN conjunct1,rule_format]
-    have ifif: "\<And>k t. (\<lambda>x. if x \<in> t then if x \<in> s then f k x else 0 else 0) =
-      (\<lambda>x. if x \<in> t \<inter> s then f k x else 0)"
-      by (rule ext) auto
-    have int': "\<And>k a b. f k integrable_on cbox a b \<inter> s"
-      apply (subst integrable_restrict_univ[symmetric])
-      apply (subst ifif[symmetric])
-      apply (subst integrable_restrict_univ)
-      apply (rule int)
-      done
-    have "\<And>a b. (\<lambda>x. if x \<in> s then g x else 0) integrable_on cbox a b \<and>
-      ((\<lambda>k. integral (cbox a b) (\<lambda>x. if x \<in> s then f k x else 0)) \<longlongrightarrow>
-      integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0)) sequentially"
-    proof (rule monotone_convergence_interval, safe, goal_cases)
-      case 1
-      show ?case by (rule int)
-    next
-      case (2 _ _ _ x)
-      then show ?case
-        apply (cases "x \<in> s")
-        using assms(3)
-        apply auto
-        done
-    next
-      case (3 _ _ x)
-      then show ?case
-        apply (cases "x \<in> s")
-        using assms(4)
-        apply auto
-        done
-    next
-      case (4 a b)
-      note * = integral_nonneg
-      have "\<And>k. norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f k x else 0)) \<le> norm (integral s (f k))"
-        unfolding real_norm_def
-        apply (subst abs_of_nonneg)
-        apply (rule *[OF int])
-        apply safe
-        apply (case_tac "x \<in> s")
-        apply (drule assms(1))
-        prefer 3
-        apply (subst abs_of_nonneg)
-        apply (rule *[OF assms(2) that(1)[THEN spec]])
-        apply (subst integral_restrict_univ[symmetric,OF int])
-        unfolding ifif
-        unfolding integral_restrict_univ[OF int']
-        apply (rule integral_subset_le[OF _ int' assms(2)])
-        using assms(1)
-        apply auto
-        done
-      then show ?case
-        using assms(5)
-        unfolding bounded_iff
-        apply safe
-        apply (rule_tac x=aa in exI)
-        apply safe
-        apply (erule_tac x="integral s (f k)" in ballE)
-        apply (rule order_trans)
-        apply assumption
-        apply auto
-        done
-    qed
-    note g = conjunctD2[OF this]
-
-    have "(g has_integral i) s"
-      unfolding has_integral_alt'
-      apply safe
-      apply (rule g(1))
-    proof goal_cases
-      case (1 e)
-      then have "e/4>0"
-        by auto
-      from LIMSEQ_D [OF i this] guess N .. note N=this
-      note assms(2)[of N,unfolded has_integral_integral has_integral_alt'[of "f N"]]
-      from this[THEN conjunct2,rule_format,OF \<open>e/4>0\<close>] guess B .. note B=conjunctD2[OF this]
-      show ?case
-        apply rule
-        apply rule
-        apply (rule B)
-        apply safe
-      proof -
-        fix a b :: 'n
-        assume ab: "ball 0 B \<subseteq> cbox a b"
-        from \<open>e > 0\<close> have "e/2 > 0"
-          by auto
-        from LIMSEQ_D [OF g(2)[of a b] this] guess M .. note M=this
-        have **: "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f N x else 0) - i) < e/2"
-          apply (rule norm_triangle_half_l)
-          using B(2)[rule_format,OF ab] N[rule_format,of N]
-          apply -
-          defer
-          apply (subst norm_minus_commute)
-          apply auto
-          done
-        have *: "\<And>f1 f2 g. \<bar>f1 - i\<bar> < e / 2 \<longrightarrow> \<bar>f2 - g\<bar> < e / 2 \<longrightarrow>
-          f1 \<le> f2 \<longrightarrow> f2 \<le> i \<longrightarrow> \<bar>g - i\<bar> < e"
-          unfolding real_inner_1_right by arith
-        show "norm (integral (cbox a b) (\<lambda>x. if x \<in> s then g x else 0) - i) < e"
-          unfolding real_norm_def
-          apply (rule *[rule_format])
-          apply (rule **[unfolded real_norm_def])
-          apply (rule M[rule_format,of "M + N",unfolded real_norm_def])
-          apply (rule le_add1)
-          apply (rule integral_le[OF int int])
-          defer
-          apply (rule order_trans[OF _ i'[rule_format,of "M + N",unfolded real_inner_1_right]])
-        proof (safe, goal_cases)
-          case (2 x)
-          have "\<And>m. x \<in> s \<Longrightarrow> \<forall>n\<ge>m. (f m x)\<bullet>1 \<le> (f n x)\<bullet>1"
-            apply (rule transitive_stepwise_le)
-            using assms(3)
-            apply auto
-            done
-          then show ?case
-            by auto
-        next
-          case 1
-          show ?case
-            apply (subst integral_restrict_univ[symmetric,OF int])
-            unfolding ifif integral_restrict_univ[OF int']
-            apply (rule integral_subset_le[OF _ int'])
-            using assms
-            apply auto
-            done
-        qed
-      qed
-    qed
-    then show ?thesis
-      apply safe
-      defer
-      apply (drule integral_unique)
-      using i
-      apply auto
-      done
-  qed
-
-  have sub: "\<And>k. integral s (\<lambda>x. f k x - f 0 x) = integral s (f k) - integral s (f 0)"
-    apply (subst integral_diff)
-    apply (rule assms(1)[rule_format])+
-    apply rule
-    done
-  have "\<And>x m. x \<in> s \<Longrightarrow> \<forall>n\<ge>m. f m x \<le> f n x"
-    apply (rule transitive_stepwise_le)
-    using assms(2)
-    apply auto
-    done
-  note * = this[rule_format]
-  have "(\<lambda>x. g x - f 0 x) integrable_on s \<and> ((\<lambda>k. integral s (\<lambda>x. f (Suc k) x - f 0 x)) \<longlongrightarrow>
-    integral s (\<lambda>x. g x - f 0 x)) sequentially"
-    apply (rule lem)
-    apply safe
-  proof goal_cases
-    case (1 k x)
-    then show ?case
-      using *[of x 0 "Suc k"] by auto
-  next
-    case (2 k)
-    then show ?case
-      apply (rule integrable_diff)
-      using assms(1)
-      apply auto
-      done
-  next
-    case (3 k x)
-    then show ?case
-      using *[of x "Suc k" "Suc (Suc k)"] by auto
-  next
-    case (4 x)
-    then show ?case
-      apply -
-      apply (rule tendsto_diff)
-      using LIMSEQ_ignore_initial_segment[OF assms(3)[rule_format],of x 1]
-      apply auto
-      done
-  next
-    case 5
-    then show ?case
-      using assms(4)
-      unfolding bounded_iff
-      apply safe
-      apply (rule_tac x="a + norm (integral s (\<lambda>x. f 0 x))" in exI)
-      apply safe
-      apply (erule_tac x="integral s (\<lambda>x. f (Suc k) x)" in ballE)
-      unfolding sub
-      apply (rule order_trans[OF norm_triangle_ineq4])
-      apply auto
-      done
-  qed
-  note conjunctD2[OF this]
-  note tendsto_add[OF this(2) tendsto_const[of "integral s (f 0)"]]
-    integrable_add[OF this(1) assms(1)[rule_format,of 0]]
-  then show ?thesis
-    unfolding sub
-    apply -
-    apply rule
-    defer
-    apply (subst(asm) integral_diff)
-    using assms(1)
-    apply auto
-    apply (rule LIMSEQ_imp_Suc)
-    apply assumption
-    done
-qed
-
-lemma has_integral_monotone_convergence_increasing:
-  fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow> real"
-  assumes f: "\<And>k. (f k has_integral x k) s"
-  assumes "\<And>k x. x \<in> s \<Longrightarrow> f k x \<le> f (Suc k) x"
-  assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
-  assumes "x \<longlonglongrightarrow> x'"
-  shows "(g has_integral x') s"
-proof -
-  have x_eq: "x = (\<lambda>i. integral s (f i))"
-    by (simp add: integral_unique[OF f])
-  then have x: "{integral s (f k) |k. True} = range x"
-    by auto
-
-  have *: "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g"
-  proof (intro monotone_convergence_increasing allI ballI assms)
-    show "bounded {integral s (f k) |k. True}"
-      unfolding x by (rule convergent_imp_bounded) fact
-  qed (auto intro: f)
-  then have "integral s g = x'"
-    by (intro LIMSEQ_unique[OF _ \<open>x \<longlonglongrightarrow> x'\<close>]) (simp add: x_eq)
-  with * show ?thesis
-    by (simp add: has_integral_integral)
-qed
-
-lemma monotone_convergence_decreasing:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
-  assumes "\<forall>k. (f k) integrable_on s"
-    and "\<forall>k. \<forall>x\<in>s. f (Suc k) x \<le> f k x"
-    and "\<forall>x\<in>s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-    and "bounded {integral s (f k)| k. True}"
-  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-proof -
-  note assm = assms[rule_format]
-  have *: "{integral s (\<lambda>x. - f k x) |k. True} = op *\<^sub>R (- 1) ` {integral s (f k)| k. True}"
-    apply safe
-    unfolding image_iff
-    apply (rule_tac x="integral s (f k)" in bexI)
-    prefer 3
-    apply (rule_tac x=k in exI)
-    apply auto
-    done
-  have "(\<lambda>x. - g x) integrable_on s \<and>
-    ((\<lambda>k. integral s (\<lambda>x. - f k x)) \<longlongrightarrow> integral s (\<lambda>x. - g x)) sequentially"
-    apply (rule monotone_convergence_increasing)
-    apply safe
-    apply (rule integrable_neg)
-    apply (rule assm)
-    defer
-    apply (rule tendsto_minus)
-    apply (rule assm)
-    apply assumption
-    unfolding *
-    apply (rule bounded_scaling)
-    using assm
-    apply auto
-    done
-  note * = conjunctD2[OF this]
-  show ?thesis
-    using integrable_neg[OF *(1)] tendsto_minus[OF *(2)]
-    by auto
-qed
-
-
-subsection \<open>Absolute integrability (this is the same as Lebesgue integrability)\<close>
-
-definition absolutely_integrable_on (infixr "absolutely'_integrable'_on" 46)
-  where "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and> (\<lambda>x. (norm(f x))) integrable_on s"
-
-lemma absolutely_integrable_onI[intro?]:
-  "f integrable_on s \<Longrightarrow>
-    (\<lambda>x. (norm(f x))) integrable_on s \<Longrightarrow> f absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def
-  by auto
-
-lemma absolutely_integrable_onD[dest]:
-  assumes "f absolutely_integrable_on s"
-  shows "f integrable_on s"
-    and "(\<lambda>x. norm (f x)) integrable_on s"
-  using assms
-  unfolding absolutely_integrable_on_def
-  by auto
-
-(*lemma absolutely_integrable_on_trans[simp]:
-  fixes f::"'n::euclidean_space \<Rightarrow> real"
-  shows "(vec1 \<circ> f) absolutely_integrable_on s \<longleftrightarrow> f absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def o_def by auto*)
-
-lemma integral_norm_bound_integral:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f integrable_on s"
-    and "g integrable_on s"
-    and "\<forall>x\<in>s. norm (f x) \<le> g x"
-  shows "norm (integral s f) \<le> integral s g"
-proof -
-  have *: "\<And>x y. (\<forall>e::real. 0 < e \<longrightarrow> x < y + e) \<Longrightarrow> x \<le> y"
-    apply (rule ccontr)
-    apply (erule_tac x="x - y" in allE)
-    apply auto
-    done
-  have norm: "norm ig < dia + e"
-    if "norm sg \<le> dsa"
-    and "\<bar>dsa - dia\<bar> < e / 2"
-    and "norm (sg - ig) < e / 2"
-    for e dsa dia and sg ig :: 'a
-    apply (rule le_less_trans[OF norm_triangle_sub[of ig sg]])
-    apply (subst real_sum_of_halves[of e,symmetric])
-    unfolding add.assoc[symmetric]
-    apply (rule add_le_less_mono)
-    defer
-    apply (subst norm_minus_commute)
-    apply (rule that(3))
-    apply (rule order_trans[OF that(1)])
-    using that(2)
-    apply arith
-    done
-  have lem: "norm (integral(cbox a b) f) \<le> integral (cbox a b) g"
-    if "f integrable_on cbox a b"
-    and "g integrable_on cbox a b"
-    and "\<forall>x\<in>cbox a b. norm (f x) \<le> g x"
-    for f :: "'n \<Rightarrow> 'a" and g a b
-  proof (rule *[rule_format])
-    fix e :: real
-    assume "e > 0"
-    then have *: "e/2 > 0"
-      by auto
-    from integrable_integral[OF that(1),unfolded has_integral[of f],rule_format,OF *]
-    guess d1 .. note d1 = conjunctD2[OF this,rule_format]
-    from integrable_integral[OF that(2),unfolded has_integral[of g],rule_format,OF *]
-    guess d2 .. note d2 = conjunctD2[OF this,rule_format]
-    note gauge_inter[OF d1(1) d2(1)]
-    from fine_division_exists[OF this, of a b] guess p . note p=this
-    show "norm (integral (cbox a b) f) < integral (cbox a b) g + e"
-      apply (rule norm)
-      defer
-      apply (rule d2(2)[OF conjI[OF p(1)],unfolded real_norm_def])
-      defer
-      apply (rule d1(2)[OF conjI[OF p(1)]])
-      defer
-      apply (rule setsum_norm_le)
-    proof safe
-      fix x k
-      assume "(x, k) \<in> p"
-      note as = tagged_division_ofD(2-4)[OF p(1) this]
-      from this(3) guess u v by (elim exE) note uv=this
-      show "norm (content k *\<^sub>R f x) \<le> content k *\<^sub>R g x"
-        unfolding uv norm_scaleR
-        unfolding abs_of_nonneg[OF content_pos_le] real_scaleR_def
-        apply (rule mult_left_mono)
-        using that(3) as
-        apply auto
-        done
-    qed (insert p[unfolded fine_inter], auto)
-  qed
-
-  { presume "\<And>e. 0 < e \<Longrightarrow> norm (integral s f) < integral s g + e"
-    then show ?thesis by (rule *[rule_format]) auto }
-  fix e :: real
-  assume "e > 0"
-  then have e: "e/2 > 0"
-    by auto
-  note assms(1)[unfolded integrable_alt[of f]] note f=this[THEN conjunct1,rule_format]
-  note assms(2)[unfolded integrable_alt[of g]] note g=this[THEN conjunct1,rule_format]
-  from integrable_integral[OF assms(1),unfolded has_integral'[of f],rule_format,OF e]
-  guess B1 .. note B1=conjunctD2[OF this[rule_format],rule_format]
-  from integrable_integral[OF assms(2),unfolded has_integral'[of g],rule_format,OF e]
-  guess B2 .. note B2=conjunctD2[OF this[rule_format],rule_format]
-  from bounded_subset_cbox[OF bounded_ball, of "0::'n" "max B1 B2"]
-  guess a b by (elim exE) note ab=this[unfolded ball_max_Un]
-
-  have "ball 0 B1 \<subseteq> cbox a b"
-    using ab by auto
-  from B1(2)[OF this] guess z .. note z=conjunctD2[OF this]
-  have "ball 0 B2 \<subseteq> cbox a b"
-    using ab by auto
-  from B2(2)[OF this] guess w .. note w=conjunctD2[OF this]
-
-  show "norm (integral s f) < integral s g + e"
-    apply (rule norm)
-    apply (rule lem[OF f g, of a b])
-    unfolding integral_unique[OF z(1)] integral_unique[OF w(1)]
-    defer
-    apply (rule w(2)[unfolded real_norm_def])
-    apply (rule z(2))
-    apply safe
-    apply (case_tac "x \<in> s")
-    unfolding if_P
-    apply (rule assms(3)[rule_format])
-    apply auto
-    done
-qed
-
-lemma integral_norm_bound_integral_component:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
-  assumes "f integrable_on s"
-    and "g integrable_on s"
-    and "\<forall>x\<in>s. norm(f x) \<le> (g x)\<bullet>k"
-  shows "norm (integral s f) \<le> (integral s g)\<bullet>k"
-proof -
-  have "norm (integral s f) \<le> integral s ((\<lambda>x. x \<bullet> k) \<circ> g)"
-    apply (rule integral_norm_bound_integral[OF assms(1)])
-    apply (rule integrable_linear[OF assms(2)])
-    apply rule
-    unfolding o_def
-    apply (rule assms)
-    done
-  then show ?thesis
-    unfolding o_def integral_component_eq[OF assms(2)] .
-qed
-
-lemma has_integral_norm_bound_integral_component:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  fixes g :: "'n \<Rightarrow> 'b::euclidean_space"
-  assumes "(f has_integral i) s"
-    and "(g has_integral j) s"
-    and "\<forall>x\<in>s. norm (f x) \<le> (g x)\<bullet>k"
-  shows "norm i \<le> j\<bullet>k"
-  using integral_norm_bound_integral_component[of f s g k]
-  unfolding integral_unique[OF assms(1)] integral_unique[OF assms(2)]
-  using assms
-  by auto
-
-lemma absolutely_integrable_le:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f absolutely_integrable_on s"
-  shows "norm (integral s f) \<le> integral s (\<lambda>x. norm (f x))"
-  apply (rule integral_norm_bound_integral)
-  using assms
-  apply auto
-  done
-
-lemma absolutely_integrable_0[intro]:
-  "(\<lambda>x. 0) absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def
-  by auto
-
-lemma absolutely_integrable_cmul[intro]:
-  "f absolutely_integrable_on s \<Longrightarrow>
-    (\<lambda>x. c *\<^sub>R f x) absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def
-  using integrable_cmul[of f s c]
-  using integrable_cmul[of "\<lambda>x. norm (f x)" s "\<bar>c\<bar>"]
-  by auto
-
-lemma absolutely_integrable_neg[intro]:
-  "f absolutely_integrable_on s \<Longrightarrow>
-    (\<lambda>x. -f(x)) absolutely_integrable_on s"
-  apply (drule absolutely_integrable_cmul[where c="-1"])
-  apply auto
-  done
-
-lemma absolutely_integrable_norm[intro]:
-  "f absolutely_integrable_on s \<Longrightarrow>
-    (\<lambda>x. norm (f x)) absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def
-  by auto
-
-lemma absolutely_integrable_abs[intro]:
-  "f absolutely_integrable_on s \<Longrightarrow>
-    (\<lambda>x. \<bar>f x::real\<bar>) absolutely_integrable_on s"
-  apply (drule absolutely_integrable_norm)
-  unfolding real_norm_def
-  apply assumption
-  done
-
-lemma absolutely_integrable_on_subinterval:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  shows "f absolutely_integrable_on s \<Longrightarrow>
-    cbox a b \<subseteq> s \<Longrightarrow> f absolutely_integrable_on cbox a b"
-  unfolding absolutely_integrable_on_def
-  by (metis integrable_on_subcbox)
-
-lemma absolutely_integrable_bounded_variation:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach"
-  assumes "f absolutely_integrable_on UNIV"
-  obtains B where "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
-  apply (rule that[of "integral UNIV (\<lambda>x. norm (f x))"])
-  apply safe
-proof goal_cases
-  case prems: (1 d)
-  note d = division_ofD[OF prems(2)]
-  have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. integral i (\<lambda>x. norm (f x)))"
-    apply (rule setsum_mono,rule absolutely_integrable_le)
-    apply (drule d(4))
-    apply safe
-    apply (rule absolutely_integrable_on_subinterval[OF assms])
-    apply auto
-    done
-  also have "\<dots> \<le> integral (\<Union>d) (\<lambda>x. norm (f x))"
-    apply (subst integral_combine_division_topdown[OF _ prems(2)])
-    using integrable_on_subdivision[OF prems(2)]
-    using assms
-    apply auto
-    done
-  also have "\<dots> \<le> integral UNIV (\<lambda>x. norm (f x))"
-    apply (rule integral_subset_le)
-    using integrable_on_subdivision[OF prems(2)]
-    using assms
-    apply auto
-    done
-  finally show ?case .
-qed
-
-lemma helplemma:
-  assumes "setsum (\<lambda>x. norm (f x - g x)) s < e"
-    and "finite s"
-  shows "\<bar>setsum (\<lambda>x. norm(f x)) s - setsum (\<lambda>x. norm(g x)) s\<bar> < e"
-  unfolding setsum_subtractf[symmetric]
-  apply (rule le_less_trans[OF setsum_abs])
-  apply (rule le_less_trans[OF _ assms(1)])
-  apply (rule setsum_mono)
-  apply (rule norm_triangle_ineq3)
-  done
-
-lemma bounded_variation_absolutely_integrable_interval:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes f: "f integrable_on cbox a b"
-    and *: "\<forall>d. d division_of (cbox a b) \<longrightarrow> setsum (\<lambda>k. norm(integral k f)) d \<le> B"
-  shows "f absolutely_integrable_on cbox a b"
-proof -
-  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of (cbox a b)}"
-  have D_1: "?D \<noteq> {}"
-    by (rule elementary_interval[of a b]) auto
-  have D_2: "bdd_above (?f`?D)"
-    by (metis * mem_Collect_eq bdd_aboveI2)
-  note D = D_1 D_2
-  let ?S = "SUP x:?D. ?f x"
-  show ?thesis
-    apply (rule absolutely_integrable_onI [OF f has_integral_integrable])
-    apply (subst has_integral[of _ ?S])
-    apply safe
-  proof goal_cases
-    case e: (1 e)
-    then have "?S - e / 2 < ?S" by simp
-    then obtain d where d: "d division_of (cbox a b)" "?S - e / 2 < (\<Sum>k\<in>d. norm (integral k f))"
-      unfolding less_cSUP_iff[OF D] by auto
-    note d' = division_ofD[OF this(1)]
-
-    have "\<forall>x. \<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
-    proof
-      fix x
-      have "\<exists>da>0. \<forall>xa\<in>\<Union>{i \<in> d. x \<notin> i}. da \<le> dist x xa"
-        apply (rule separate_point_closed)
-        apply (rule closed_Union)
-        apply (rule finite_subset[OF _ d'(1)])
-        using d'(4)
-        apply auto
-        done
-      then show "\<exists>e>0. \<forall>i\<in>d. x \<notin> i \<longrightarrow> ball x e \<inter> i = {}"
-        by force
-    qed
-    from choice[OF this] guess k .. note k=conjunctD2[OF this[rule_format],rule_format]
-
-    have "e/2 > 0"
-      using e by auto
-    from henstock_lemma[OF assms(1) this] guess g . note g=this[rule_format]
-    let ?g = "\<lambda>x. g x \<inter> ball x (k x)"
-    show ?case
-      apply (rule_tac x="?g" in exI)
-      apply safe
-    proof -
-      show "gauge ?g"
-        using g(1) k(1)
-        unfolding gauge_def
-        by auto
-      fix p
-      assume "p tagged_division_of (cbox a b)" and "?g fine p"
-      note p = this(1) conjunctD2[OF this(2)[unfolded fine_inter]]
-      note p' = tagged_division_ofD[OF p(1)]
-      define p' where "p' = {(x,k) | x k. \<exists>i l. x \<in> i \<and> i \<in> d \<and> (x,l) \<in> p \<and> k = i \<inter> l}"
-      have gp': "g fine p'"
-        using p(2)
-        unfolding p'_def fine_def
-        by auto
-      have p'': "p' tagged_division_of (cbox a b)"
-        apply (rule tagged_division_ofI)
-      proof -
-        show "finite p'"
-          apply (rule finite_subset[of _ "(\<lambda>(k,(x,l)). (x,k \<inter> l)) `
-            {(k,xl) | k xl. k \<in> d \<and> xl \<in> p}"])
-          unfolding p'_def
-          defer
-          apply (rule finite_imageI,rule finite_product_dependent[OF d'(1) p'(1)])
-          apply safe
-          unfolding image_iff
-          apply (rule_tac x="(i,x,l)" in bexI)
-          apply auto
-          done
-        fix x k
-        assume "(x, k) \<in> p'"
-        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
-          unfolding p'_def by auto
-        then guess i l by (elim exE) note il=conjunctD4[OF this]
-        show "x \<in> k" and "k \<subseteq> cbox a b"
-          using p'(2-3)[OF il(3)] il by auto
-        show "\<exists>a b. k = cbox a b"
-          unfolding il using p'(4)[OF il(3)] d'(4)[OF il(2)]
-          apply safe
-          unfolding inter_interval
-          apply auto
-          done
-      next
-        fix x1 k1
-        assume "(x1, k1) \<in> p'"
-        then have "\<exists>i l. x1 \<in> i \<and> i \<in> d \<and> (x1, l) \<in> p \<and> k1 = i \<inter> l"
-          unfolding p'_def by auto
-        then guess i1 l1 by (elim exE) note il1=conjunctD4[OF this]
-        fix x2 k2
-        assume "(x2,k2)\<in>p'"
-        then have "\<exists>i l. x2 \<in> i \<and> i \<in> d \<and> (x2, l) \<in> p \<and> k2 = i \<inter> l"
-          unfolding p'_def by auto
-        then guess i2 l2 by (elim exE) note il2=conjunctD4[OF this]
-        assume "(x1, k1) \<noteq> (x2, k2)"
-        then have "interior i1 \<inter> interior i2 = {} \<or> interior l1 \<inter> interior l2 = {}"
-          using d'(5)[OF il1(2) il2(2)] p'(5)[OF il1(3) il2(3)]
-          unfolding il1 il2
-          by auto
-        then show "interior k1 \<inter> interior k2 = {}"
-          unfolding il1 il2 by auto
-      next
-        have *: "\<forall>(x, X) \<in> p'. X \<subseteq> cbox a b"
-          unfolding p'_def using d' by auto
-        show "\<Union>{k. \<exists>x. (x, k) \<in> p'} = cbox a b"
-          apply rule
-          apply (rule Union_least)
-          unfolding mem_Collect_eq
-          apply (erule exE)
-          apply (drule *[rule_format])
-          apply safe
-        proof -
-          fix y
-          assume y: "y \<in> cbox a b"
-          then have "\<exists>x l. (x, l) \<in> p \<and> y\<in>l"
-            unfolding p'(6)[symmetric] by auto
-          then guess x l by (elim exE) note xl=conjunctD2[OF this]
-          then have "\<exists>k. k \<in> d \<and> y \<in> k"
-            using y unfolding d'(6)[symmetric] by auto
-          then guess i .. note i = conjunctD2[OF this]
-          have "x \<in> i"
-            using fineD[OF p(3) xl(1)]
-            using k(2)[OF i(1), of x]
-            using i(2) xl(2)
-            by auto
-          then show "y \<in> \<Union>{k. \<exists>x. (x, k) \<in> p'}"
-            unfolding p'_def Union_iff
-            apply (rule_tac x="i \<inter> l" in bexI)
-            using i xl
-            apply auto
-            done
-        qed
-      qed
-
-      then have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x - integral k f)) < e / 2"
-        apply -
-        apply (rule g(2)[rule_format])
-        unfolding tagged_division_of_def
-        apply safe
-        apply (rule gp')
-        done
-      then have **: "\<bar>(\<Sum>(x,k)\<in>p'. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p'. norm (integral k f))\<bar> < e / 2"
-        unfolding split_def
-        using p''
-        by (force intro!: helplemma)
-
-      have p'alt: "p' = {(x,(i \<inter> l)) | x i l. (x,l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}}"
-      proof (safe, goal_cases)
-        case prems: (2 _ _ x i l)
-        have "x \<in> i"
-          using fineD[OF p(3) prems(1)] k(2)[OF prems(2), of x] prems(4-)
-          by auto
-        then have "(x, i \<inter> l) \<in> p'"
-          unfolding p'_def
-          using prems
-          apply safe
-          apply (rule_tac x=x in exI)
-          apply (rule_tac x="i \<inter> l" in exI)
-          apply safe
-          using prems
-          apply auto
-          done
-        then show ?case
-          using prems(3) by auto
-      next
-        fix x k
-        assume "(x, k) \<in> p'"
-        then have "\<exists>i l. x \<in> i \<and> i \<in> d \<and> (x, l) \<in> p \<and> k = i \<inter> l"
-          unfolding p'_def by auto
-        then guess i l by (elim exE) note il=conjunctD4[OF this]
-        then show "\<exists>y i l. (x, k) = (y, i \<inter> l) \<and> (y, l) \<in> p \<and> i \<in> d \<and> i \<inter> l \<noteq> {}"
-          apply (rule_tac x=x in exI)
-          apply (rule_tac x=i in exI)
-          apply (rule_tac x=l in exI)
-          using p'(2)[OF il(3)]
-          apply auto
-          done
-      qed
-      have sum_p': "(\<Sum>(x, k)\<in>p'. norm (integral k f)) = (\<Sum>k\<in>snd ` p'. norm (integral k f))"
-        apply (subst setsum.over_tagged_division_lemma[OF p'',of "\<lambda>k. norm (integral k f)"])
-        unfolding norm_eq_zero
-        apply (rule integral_null)
-        apply assumption
-        apply rule
-        done
-      note snd_p = division_ofD[OF division_of_tagged_division[OF p(1)]]
-
-      have *: "\<And>sni sni' sf sf'. \<bar>sf' - sni'\<bar> < e / 2 \<longrightarrow> ?S - e / 2 < sni \<and> sni' \<le> ?S \<and>
-        sni \<le> sni' \<and> sf' = sf \<longrightarrow> \<bar>sf - ?S\<bar> < e"
-        by arith
-      show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - ?S) < e"
-        unfolding real_norm_def
-        apply (rule *[rule_format,OF **])
-        apply safe
-        apply(rule d(2))
-      proof goal_cases
-        case 1
-        show ?case
-          by (auto simp: sum_p' division_of_tagged_division[OF p''] D intro!: cSUP_upper)
-      next
-        case 2
-        have *: "{k \<inter> l | k l. k \<in> d \<and> l \<in> snd ` p} =
-          (\<lambda>(k,l). k \<inter> l) ` {(k,l)|k l. k \<in> d \<and> l \<in> snd ` p}"
-          by auto
-        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>i\<in>d. \<Sum>l\<in>snd ` p. norm (integral (i \<inter> l) f))"
-        proof (rule setsum_mono, goal_cases)
-          case k: (1 k)
-          from d'(4)[OF this] guess u v by (elim exE) note uv=this
-          define d' where "d' = {cbox u v \<inter> l |l. l \<in> snd ` p \<and>  cbox u v \<inter> l \<noteq> {}}"
-          note uvab = d'(2)[OF k[unfolded uv]]
-          have "d' division_of cbox u v"
-            apply (subst d'_def)
-            apply (rule division_inter_1)
-            apply (rule division_of_tagged_division[OF p(1)])
-            apply (rule uvab)
-            done
-          then have "norm (integral k f) \<le> setsum (\<lambda>k. norm (integral k f)) d'"
-            unfolding uv
-            apply (subst integral_combine_division_topdown[of _ _ d'])
-            apply (rule integrable_on_subcbox[OF assms(1) uvab])
-            apply assumption
-            apply (rule setsum_norm_le)
-            apply auto
-            done
-          also have "\<dots> = (\<Sum>k\<in>{k \<inter> l |l. l \<in> snd ` p}. norm (integral k f))"
-            apply (rule setsum.mono_neutral_left)
-            apply (subst simple_image)
-            apply (rule finite_imageI)+
-            apply fact
-            unfolding d'_def uv
-            apply blast
-          proof (rule, goal_cases)
-            case prems: (1 i)
-            then have "i \<in> {cbox u v \<inter> l |l. l \<in> snd ` p}"
-              by auto
-            from this[unfolded mem_Collect_eq] guess l .. note l=this
-            then have "cbox u v \<inter> l = {}"
-              using prems by auto
-            then show ?case
-              using l by auto
-          qed
-          also have "\<dots> = (\<Sum>l\<in>snd ` p. norm (integral (k \<inter> l) f))"
-            unfolding simple_image
-            apply (rule setsum.reindex_nontrivial [unfolded o_def])
-            apply (rule finite_imageI)
-            apply (rule p')
-          proof goal_cases
-            case prems: (1 l y)
-            have "interior (k \<inter> l) \<subseteq> interior (l \<inter> y)"
-              apply (subst(2) interior_Int)
-              apply (rule Int_greatest)
-              defer
-              apply (subst prems(4))
-              apply auto
-              done
-            then have *: "interior (k \<inter> l) = {}"
-              using snd_p(5)[OF prems(1-3)] by auto
-            from d'(4)[OF k] snd_p(4)[OF prems(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
-            show ?case
-              using *
-              unfolding uv inter_interval content_eq_0_interior[symmetric]
-              by auto
-          qed
-          finally show ?case .
-        qed
-        also have "\<dots> = (\<Sum>(i,l)\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (i\<inter>l) f))"
-          apply (subst sum_sum_product[symmetric])
-          apply fact
-          using p'(1)
-          apply auto
-          done
-        also have "\<dots> = (\<Sum>x\<in>{(i, l) |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral (case_prod op \<inter> x) f))"
-          unfolding split_def ..
-        also have "\<dots> = (\<Sum>k\<in>{i \<inter> l |i l. i \<in> d \<and> l \<in> snd ` p}. norm (integral k f))"
-          unfolding *
-          apply (rule setsum.reindex_nontrivial [symmetric, unfolded o_def])
-          apply (rule finite_product_dependent)
-          apply fact
-          apply (rule finite_imageI)
-          apply (rule p')
-          unfolding split_paired_all mem_Collect_eq split_conv o_def
-        proof -
-          note * = division_ofD(4,5)[OF division_of_tagged_division,OF p(1)]
-          fix l1 l2 k1 k2
-          assume as:
-            "(l1, k1) \<noteq> (l2, k2)"
-            "l1 \<inter> k1 = l2 \<inter> k2"
-            "\<exists>i l. (l1, k1) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
-            "\<exists>i l. (l2, k2) = (i, l) \<and> i \<in> d \<and> l \<in> snd ` p"
-          then have "l1 \<in> d" and "k1 \<in> snd ` p"
-            by auto from d'(4)[OF this(1)] *(1)[OF this(2)]
-          guess u1 v1 u2 v2 by (elim exE) note uv=this
-          have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
-            using as by auto
-          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
-            apply -
-            apply (erule disjE)
-            apply (rule disjI2)
-            apply (rule d'(5))
-            prefer 4
-            apply (rule disjI1)
-            apply (rule *)
-            using as
-            apply auto
-            done
-          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
-            using as(2) by auto
-          ultimately have "interior(l1 \<inter> k1) = {}"
-            by auto
-          then show "norm (integral (l1 \<inter> k1) f) = 0"
-            unfolding uv inter_interval
-            unfolding content_eq_0_interior[symmetric]
-            by auto
-        qed
-        also have "\<dots> = (\<Sum>(x, k)\<in>p'. norm (integral k f))"
-          unfolding sum_p'
-          apply (rule setsum.mono_neutral_right)
-          apply (subst *)
-          apply (rule finite_imageI[OF finite_product_dependent])
-          apply fact
-          apply (rule finite_imageI[OF p'(1)])
-          apply safe
-        proof goal_cases
-          case (2 i ia l a b)
-          then have "ia \<inter> b = {}"
-            unfolding p'alt image_iff Bex_def not_ex
-            apply (erule_tac x="(a, ia \<inter> b)" in allE)
-            apply auto
-            done
-          then show ?case
-            by auto
-        next
-          case (1 x a b)
-          then show ?case
-            unfolding p'_def
-            apply safe
-            apply (rule_tac x=i in exI)
-            apply (rule_tac x=l in exI)
-            unfolding snd_conv image_iff
-            apply safe
-            apply (rule_tac x="(a,l)" in bexI)
-            apply auto
-            done
-        qed
-        finally show ?case .
-      next
-        case 3
-        let ?S = "{(x, i \<inter> l) |x i l. (x, l) \<in> p \<and> i \<in> d}"
-        have Sigma_alt: "\<And>s t. s \<times> t = {(i, j) |i j. i \<in> s \<and> j \<in> t}"
-          by auto
-        have *: "?S = (\<lambda>(xl,i). (fst xl, snd xl \<inter> i)) ` (p \<times> d)"
-          apply safe
-          unfolding image_iff
-          apply (rule_tac x="((x,l),i)" in bexI)
-          apply auto
-          done
-        note pdfin = finite_cartesian_product[OF p'(1) d'(1)]
-        have "(\<Sum>(x, k)\<in>p'. norm (content k *\<^sub>R f x)) = (\<Sum>(x, k)\<in>?S. \<bar>content k\<bar> * norm (f x))"
-          unfolding norm_scaleR
-          apply (rule setsum.mono_neutral_left)
-          apply (subst *)
-          apply (rule finite_imageI)
-          apply fact
-          unfolding p'alt
-          apply blast
-          apply safe
-          apply (rule_tac x=x in exI)
-          apply (rule_tac x=i in exI)
-          apply (rule_tac x=l in exI)
-          apply auto
-          done
-        also have "\<dots> = (\<Sum>((x,l),i)\<in>p \<times> d. \<bar>content (l \<inter> i)\<bar> * norm (f x))"
-          unfolding *
-          apply (subst setsum.reindex_nontrivial)
-          apply fact
-          unfolding split_paired_all
-          unfolding o_def split_def snd_conv fst_conv mem_Sigma_iff prod.inject
-          apply (elim conjE)
-        proof -
-          fix x1 l1 k1 x2 l2 k2
-          assume as: "(x1, l1) \<in> p" "(x2, l2) \<in> p" "k1 \<in> d" "k2 \<in> d"
-            "x1 = x2" "l1 \<inter> k1 = l2 \<inter> k2" "\<not> ((x1 = x2 \<and> l1 = l2) \<and> k1 = k2)"
-          from d'(4)[OF as(3)] p'(4)[OF as(1)] guess u1 v1 u2 v2 by (elim exE) note uv=this
-          from as have "l1 \<noteq> l2 \<or> k1 \<noteq> k2"
-            by auto
-          then have "interior k1 \<inter> interior k2 = {} \<or> interior l1 \<inter> interior l2 = {}"
-            apply -
-            apply (erule disjE)
-            apply (rule disjI2)
-            defer
-            apply (rule disjI1)
-            apply (rule d'(5)[OF as(3-4)])
-            apply assumption
-            apply (rule p'(5)[OF as(1-2)])
-            apply auto
-            done
-          moreover have "interior (l1 \<inter> k1) = interior (l2 \<inter> k2)"
-            unfolding  as ..
-          ultimately have "interior (l1 \<inter> k1) = {}"
-            by auto
-          then show "\<bar>content (l1 \<inter> k1)\<bar> * norm (f x1) = 0"
-            unfolding uv inter_interval
-            unfolding content_eq_0_interior[symmetric]
-            by auto
-        qed safe
-        also have "\<dots> = (\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x))"
-          unfolding Sigma_alt
-          apply (subst sum_sum_product[symmetric])
-          apply (rule p')
-          apply rule
-          apply (rule d')
-          apply (rule setsum.cong)
-          apply (rule refl)
-          unfolding split_paired_all split_conv
-        proof -
-          fix x l
-          assume as: "(x, l) \<in> p"
-          note xl = p'(2-4)[OF this]
-          from this(3) guess u v by (elim exE) note uv=this
-          have "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar>) = (\<Sum>k\<in>d. content (k \<inter> cbox u v))"
-            apply (rule setsum.cong)
-            apply (rule refl)
-            apply (drule d'(4))
-            apply safe
-            apply (subst Int_commute)
-            unfolding inter_interval uv
-            apply (subst abs_of_nonneg)
-            apply auto
-            done
-          also have "\<dots> = setsum content {k \<inter> cbox u v| k. k \<in> d}"
-            unfolding simple_image
-            apply (rule setsum.reindex_nontrivial [unfolded o_def, symmetric])
-            apply (rule d')
-          proof goal_cases
-            case prems: (1 k y)
-            from d'(4)[OF this(1)] d'(4)[OF this(2)]
-            guess u1 v1 u2 v2 by (elim exE) note uv=this
-            have "{} = interior ((k \<inter> y) \<inter> cbox u v)"
-              apply (subst interior_Int)
-              using d'(5)[OF prems(1-3)]
-              apply auto
-              done
-            also have "\<dots> = interior (y \<inter> (k \<inter> cbox u v))"
-              by auto
-            also have "\<dots> = interior (k \<inter> cbox u v)"
-              unfolding prems(4) by auto
-            finally show ?case
-              unfolding uv inter_interval content_eq_0_interior ..
-          qed
-          also have "\<dots> = setsum content {cbox u v \<inter> k |k. k \<in> d \<and> cbox u v \<inter> k \<noteq> {}}"
-            apply (rule setsum.mono_neutral_right)
-            unfolding simple_image
-            apply (rule finite_imageI)
-            apply (rule d')
-            apply blast
-            apply safe
-            apply (rule_tac x=k in exI)
-          proof goal_cases
-            case prems: (1 i k)
-            from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this
-            have "interior (k \<inter> cbox u v) \<noteq> {}"
-              using prems(2)
-              unfolding ab inter_interval content_eq_0_interior
-              by auto
-            then show ?case
-              using prems(1)
-              using interior_subset[of "k \<inter> cbox u v"]
-              by auto
-          qed
-          finally show "(\<Sum>i\<in>d. \<bar>content (l \<inter> i)\<bar> * norm (f x)) = content l *\<^sub>R norm (f x)"
-            unfolding setsum_left_distrib[symmetric] real_scaleR_def
-            apply (subst(asm) additive_content_division[OF division_inter_1[OF d(1)]])
-            using xl(2)[unfolded uv]
-            unfolding uv
-            apply auto
-            done
-        qed
-        finally show ?case .
-      qed
-    qed
-  qed
-qed
-
-lemma bounded_variation_absolutely_integrable:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "f integrable_on UNIV"
-    and "\<forall>d. d division_of (\<Union>d) \<longrightarrow> setsum (\<lambda>k. norm (integral k f)) d \<le> B"
-  shows "f absolutely_integrable_on UNIV"
-proof (rule absolutely_integrable_onI, fact, rule)
-  let ?f = "\<lambda>d. \<Sum>k\<in>d. norm (integral k f)" and ?D = "{d. d division_of  (\<Union>d)}"
-  have D_1: "?D \<noteq> {}"
-    by (rule elementary_interval) auto
-  have D_2: "bdd_above (?f`?D)"
-    by (intro bdd_aboveI2[where M=B] assms(2)[rule_format]) simp
-  note D = D_1 D_2
-  let ?S = "SUP d:?D. ?f d"
-  have f_int: "\<And>a b. f absolutely_integrable_on cbox a b"
-    apply (rule bounded_variation_absolutely_integrable_interval[where B=B])
-    apply (rule integrable_on_subcbox[OF assms(1)])
-    defer
-    apply safe
-    apply (rule assms(2)[rule_format])
-    apply auto
-    done
-  show "((\<lambda>x. norm (f x)) has_integral ?S) UNIV"
-    apply (subst has_integral_alt')
-    apply safe
-  proof goal_cases
-    case (1 a b)
-    show ?case
-      using f_int[of a b] by auto
-  next
-    case prems: (2 e)
-    have "\<exists>y\<in>setsum (\<lambda>k. norm (integral k f)) ` {d. d division_of \<Union>d}. \<not> y \<le> ?S - e"
-    proof (rule ccontr)
-      assume "\<not> ?thesis"
-      then have "?S \<le> ?S - e"
-        by (intro cSUP_least[OF D(1)]) auto
-      then show False
-        using prems by auto
-    qed
-    then obtain K where *: "\<exists>x\<in>{d. d division_of \<Union>d}. K = (\<Sum>k\<in>x. norm (integral k f))"
-      "SUPREMUM {d. d division_of \<Union>d} (setsum (\<lambda>k. norm (integral k f))) - e < K"
-      by (auto simp add: image_iff not_le)
-    from this(1) obtain d where "d division_of \<Union>d"
-      and "K = (\<Sum>k\<in>d. norm (integral k f))"
-      by auto
-    note d = this(1) *(2)[unfolded this(2)]
-    note d'=division_ofD[OF this(1)]
-    have "bounded (\<Union>d)"
-      by (rule elementary_bounded,fact)
-    from this[unfolded bounded_pos] obtain K where
-       K: "0 < K" "\<forall>x\<in>\<Union>d. norm x \<le> K" by auto
-    show ?case
-      apply (rule_tac x="K + 1" in exI)
-      apply safe
-    proof -
-      fix a b :: 'n
-      assume ab: "ball 0 (K + 1) \<subseteq> cbox a b"
-      have *: "\<forall>s s1. ?S - e < s1 \<and> s1 \<le> s \<and> s < ?S + e \<longrightarrow> \<bar>s - ?S\<bar> < e"
-        by arith
-      show "norm (integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) - ?S) < e"
-        unfolding real_norm_def
-        apply (rule *[rule_format])
-        apply safe
-        apply (rule d(2))
-      proof goal_cases
-        case 1
-        have "(\<Sum>k\<in>d. norm (integral k f)) \<le> setsum (\<lambda>k. integral k (\<lambda>x. norm (f x))) d"
-          apply (rule setsum_mono)
-          apply (rule absolutely_integrable_le)
-          apply (drule d'(4))
-          apply safe
-          apply (rule f_int)
-          done
-        also have "\<dots> = integral (\<Union>d) (\<lambda>x. norm (f x))"
-          apply (rule integral_combine_division_bottomup[symmetric])
-          apply (rule d)
-          unfolding forall_in_division[OF d(1)]
-          using f_int
-          apply auto
-          done
-        also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0)"
-        proof -
-          have "\<Union>d \<subseteq> cbox a b"
-            apply rule
-            apply (drule K(2)[rule_format])
-            apply (rule ab[unfolded subset_eq,rule_format])
-            apply (auto simp add: dist_norm)
-            done
-          then show ?thesis
-            apply -
-            apply (subst if_P)
-            apply rule
-            apply (rule integral_subset_le)
-            defer
-            apply (rule integrable_on_subdivision[of _ _ _ "cbox a b"])
-            apply (rule d)
-            using f_int[of a b]
-            apply auto
-            done
-        qed
-        finally show ?case .
-      next
-        note f = absolutely_integrable_onD[OF f_int[of a b]]
-        note * = this(2)[unfolded has_integral_integral has_integral[of "\<lambda>x. norm (f x)"],rule_format]
-        have "e/2>0"
-          using \<open>e > 0\<close> by auto
-        from * [OF this] obtain d1 where
-          d1: "gauge d1" "\<forall>p. p tagged_division_of (cbox a b) \<and> d1 fine p \<longrightarrow>
-            norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm (f x))) < e / 2"
-          by auto
-        from henstock_lemma [OF f(1) \<open>e/2>0\<close>] obtain d2 where
-          d2: "gauge d2" "\<forall>p. p tagged_partial_division_of (cbox a b) \<and> d2 fine p \<longrightarrow>
-            (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x - integral k f)) < e / 2" .
-         obtain p where
-          p: "p tagged_division_of (cbox a b)" "d1 fine p" "d2 fine p"
-          by (rule fine_division_exists [OF gauge_inter [OF d1(1) d2(1)], of a b])
-            (auto simp add: fine_inter)
-        have *: "\<And>sf sf' si di. sf' = sf \<longrightarrow> si \<le> ?S \<longrightarrow> \<bar>sf - si\<bar> < e / 2 \<longrightarrow>
-          \<bar>sf' - di\<bar> < e / 2 \<longrightarrow> di < ?S + e"
-          by arith
-        show "integral (cbox a b) (\<lambda>x. if x \<in> UNIV then norm (f x) else 0) < ?S + e"
-          apply (subst if_P)
-          apply rule
-        proof (rule *[rule_format])
-          show "\<bar>(\<Sum>(x,k)\<in>p. norm (content k *\<^sub>R f x)) - (\<Sum>(x,k)\<in>p. norm (integral k f))\<bar> < e / 2"
-            unfolding split_def
-            apply (rule helplemma)
-            using d2(2)[rule_format,of p]
-            using p(1,3)
-            unfolding tagged_division_of_def split_def
-            apply auto
-            done
-          show "\<bar>(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) - integral (cbox a b) (\<lambda>x. norm(f x))\<bar> < e / 2"
-            using d1(2)[rule_format,OF conjI[OF p(1,2)]]
-            by (simp only: real_norm_def)
-          show "(\<Sum>(x, k)\<in>p. content k *\<^sub>R norm (f x)) = (\<Sum>(x, k)\<in>p. norm (content k *\<^sub>R f x))"
-            apply (rule setsum.cong)
-            apply (rule refl)
-            unfolding split_paired_all split_conv
-            apply (drule tagged_division_ofD(4)[OF p(1)])
-            unfolding norm_scaleR
-            apply (subst abs_of_nonneg)
-            apply auto
-            done
-          show "(\<Sum>(x, k)\<in>p. norm (integral k f)) \<le> ?S"
-            using partial_division_of_tagged_division[of p "cbox a b"] p(1)
-            apply (subst setsum.over_tagged_division_lemma[OF p(1)])
-            apply (simp add: integral_null)
-            apply (intro cSUP_upper2[OF D(2), of "snd ` p"])
-            apply (auto simp: tagged_partial_division_of_def)
-            done
-        qed
-      qed
-    qed (insert K, auto)
-  qed
-qed
-
-lemma absolutely_integrable_restrict_univ:
-  "(\<lambda>x. if x \<in> s then f x else (0::'a::banach)) absolutely_integrable_on UNIV \<longleftrightarrow>
-    f absolutely_integrable_on s"
-  unfolding absolutely_integrable_on_def if_distrib norm_zero integrable_restrict_univ ..
-
-lemma absolutely_integrable_add[intro]:
-  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "f absolutely_integrable_on s"
-    and "g absolutely_integrable_on s"
-  shows "(\<lambda>x. f x + g x) absolutely_integrable_on s"
-proof -
-  let ?P = "\<And>f g::'n \<Rightarrow> 'm. f absolutely_integrable_on UNIV \<Longrightarrow>
-    g absolutely_integrable_on UNIV \<Longrightarrow> (\<lambda>x. f x + g x) absolutely_integrable_on UNIV"
-  {
-    presume as: "PROP ?P"
-    note a = absolutely_integrable_restrict_univ[symmetric]
-    have *: "\<And>x. (if x \<in> s then f x else 0) + (if x \<in> s then g x else 0) =
-      (if x \<in> s then f x + g x else 0)" by auto
-    show ?thesis
-      apply (subst a)
-      using as[OF assms[unfolded a[of f] a[of g]]]
-      apply (simp only: *)
-      done
-  }
-  fix f g :: "'n \<Rightarrow> 'm"
-  assume assms: "f absolutely_integrable_on UNIV" "g absolutely_integrable_on UNIV"
-  note absolutely_integrable_bounded_variation
-  from this[OF assms(1)] this[OF assms(2)] guess B1 B2 . note B=this[rule_format]
-  show "(\<lambda>x. f x + g x) absolutely_integrable_on UNIV"
-    apply (rule bounded_variation_absolutely_integrable[of _ "B1+B2"])
-    apply (rule integrable_add)
-    prefer 3
-    apply safe
-  proof goal_cases
-    case prems: (1 d)
-    have "\<And>k. k \<in> d \<Longrightarrow> f integrable_on k \<and> g integrable_on k"
-      apply (drule division_ofD(4)[OF prems])
-      apply safe
-      apply (rule_tac[!] integrable_on_subcbox[of _ UNIV])
-      using assms
-      apply auto
-      done
-    then have "(\<Sum>k\<in>d. norm (integral k (\<lambda>x. f x + g x))) \<le>
-      (\<Sum>k\<in>d. norm (integral k f)) + (\<Sum>k\<in>d. norm (integral k g))"
-      apply -
-      unfolding setsum.distrib [symmetric]
-      apply (rule setsum_mono)
-      apply (subst integral_add)
-      prefer 3
-      apply (rule norm_triangle_ineq)
-      apply auto
-      done
-    also have "\<dots> \<le> B1 + B2"
-      using B(1)[OF prems] B(2)[OF prems] by auto
-    finally show ?case .
-  qed (insert assms, auto)
-qed
-
-lemma absolutely_integrable_sub[intro]:
-  fixes f g :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "f absolutely_integrable_on s"
-    and "g absolutely_integrable_on s"
-  shows "(\<lambda>x. f x - g x) absolutely_integrable_on s"
-  using absolutely_integrable_add[OF assms(1) absolutely_integrable_neg[OF assms(2)]]
-  by (simp add: algebra_simps)
-
-lemma absolutely_integrable_linear:
-  fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
-    and h :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
-  assumes "f absolutely_integrable_on s"
-    and "bounded_linear h"
-  shows "(h \<circ> f) absolutely_integrable_on s"
-proof -
-  {
-    presume as: "\<And>f::'m \<Rightarrow> 'n. \<And>h::'n \<Rightarrow> 'p. f absolutely_integrable_on UNIV \<Longrightarrow>
-      bounded_linear h \<Longrightarrow> (h \<circ> f) absolutely_integrable_on UNIV"
-    note a = absolutely_integrable_restrict_univ[symmetric]
-    show ?thesis
-      apply (subst a)
-      using as[OF assms[unfolded a[of f] a[of g]]]
-      apply (simp only: o_def if_distrib linear_simps[OF assms(2)])
-      done
-  }
-  fix f :: "'m \<Rightarrow> 'n"
-  fix h :: "'n \<Rightarrow> 'p"
-  assume assms: "f absolutely_integrable_on UNIV" "bounded_linear h"
-  from absolutely_integrable_bounded_variation[OF assms(1)] guess B . note B=this
-  from bounded_linear.pos_bounded[OF assms(2)] guess b .. note b=conjunctD2[OF this]
-  show "(h \<circ> f) absolutely_integrable_on UNIV"
-    apply (rule bounded_variation_absolutely_integrable[of _ "B * b"])
-    apply (rule integrable_linear[OF _ assms(2)])
-    apply safe
-  proof goal_cases
-    case prems: (2 d)
-    have "(\<Sum>k\<in>d. norm (integral k (h \<circ> f))) \<le> setsum (\<lambda>k. norm(integral k f)) d * b"
-      unfolding setsum_left_distrib
-      apply (rule setsum_mono)
-    proof goal_cases
-      case (1 k)
-      from division_ofD(4)[OF prems this]
-      guess u v by (elim exE) note uv=this
-      have *: "f integrable_on k"
-        unfolding uv
-        apply (rule integrable_on_subcbox[of _ UNIV])
-        using assms
-        apply auto
-        done
-      note this[unfolded has_integral_integral]
-      note has_integral_linear[OF this assms(2)] integrable_linear[OF * assms(2)]
-      note * = has_integral_unique[OF this(2)[unfolded has_integral_integral] this(1)]
-      show ?case
-        unfolding * using b by auto
-    qed
-    also have "\<dots> \<le> B * b"
-      apply (rule mult_right_mono)
-      using B prems b
-      apply auto
-      done
-    finally show ?case .
-  qed (insert assms, auto)
-qed
-
-lemma absolutely_integrable_setsum:
-  fixes f :: "'a \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "finite t"
-    and "\<And>a. a \<in> t \<Longrightarrow> (f a) absolutely_integrable_on s"
-  shows "(\<lambda>x. setsum (\<lambda>a. f a x) t) absolutely_integrable_on s"
-  using assms(1,2)
-  by induct auto
-
-lemma absolutely_integrable_vector_abs:
-  fixes f :: "'a::euclidean_space => 'b::euclidean_space"
-    and T :: "'c::euclidean_space \<Rightarrow> 'b"
-  assumes f: "f absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>T i\<bar> *\<^sub>R i)) absolutely_integrable_on s"
-  (is "?Tf absolutely_integrable_on s")
-proof -
-  have if_distrib: "\<And>P A B x. (if P then A else B) *\<^sub>R x = (if P then A *\<^sub>R x else B *\<^sub>R x)"
-    by simp
-  have *: "\<And>x. ?Tf x = (\<Sum>i\<in>Basis.
-    ((\<lambda>y. (\<Sum>j\<in>Basis. (if j = i then y else 0) *\<^sub>R j)) o
-     (\<lambda>x. (norm (\<Sum>j\<in>Basis. (if j = i then f x\<bullet>T i else 0) *\<^sub>R j)))) x)"
-    by (simp add: comp_def if_distrib setsum.If_cases)
-  show ?thesis
-    unfolding *
-    apply (rule absolutely_integrable_setsum[OF finite_Basis])
-    apply (rule absolutely_integrable_linear)
-    apply (rule absolutely_integrable_norm)
-    apply (rule absolutely_integrable_linear[OF f, unfolded o_def])
-    apply (auto simp: linear_linear euclidean_eq_iff[where 'a='c] inner_simps intro!: linearI)
-    done
-qed
-
-lemma absolutely_integrable_max:
-  fixes f g :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
-  assumes "f absolutely_integrable_on s"
-    and "g absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
-proof -
-  have *:"\<And>x. (1 / 2) *\<^sub>R (((\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i)::'n) + (f x + g x)) =
-      (\<Sum>i\<in>Basis. max (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
-    unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
-  note absolutely_integrable_sub[OF assms] absolutely_integrable_add[OF assms]
-  note absolutely_integrable_vector_abs[OF this(1), where T="\<lambda>x. x"] this(2)
-  note absolutely_integrable_add[OF this]
-  note absolutely_integrable_cmul[OF this, of "1/2"]
-  then show ?thesis unfolding * .
-qed
-
-lemma absolutely_integrable_min:
-  fixes f g::"'m::euclidean_space \<Rightarrow> 'n::euclidean_space"
-  assumes "f absolutely_integrable_on s"
-    and "g absolutely_integrable_on s"
-  shows "(\<lambda>x. (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)::'n) absolutely_integrable_on s"
-proof -
-  have *:"\<And>x. (1 / 2) *\<^sub>R ((f x + g x) - (\<Sum>i\<in>Basis. \<bar>(f x - g x) \<bullet> i\<bar> *\<^sub>R i::'n)) =
-      (\<Sum>i\<in>Basis. min (f(x)\<bullet>i) (g(x)\<bullet>i) *\<^sub>R i)"
-    unfolding euclidean_eq_iff[where 'a='n] by (auto simp: inner_simps)
-
-  note absolutely_integrable_add[OF assms] absolutely_integrable_sub[OF assms]
-  note this(1) absolutely_integrable_vector_abs[OF this(2), where T="\<lambda>x. x"]
-  note absolutely_integrable_sub[OF this]
-  note absolutely_integrable_cmul[OF this,of "1/2"]
-  then show ?thesis unfolding * .
-qed
-
-lemma absolutely_integrable_abs_eq:
-  fixes f::"'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  shows "f absolutely_integrable_on s \<longleftrightarrow> f integrable_on s \<and>
-    (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on s"
-  (is "?l = ?r")
-proof
-  assume ?l
-  then show ?r
-    apply -
-    apply rule
-    defer
-    apply (drule absolutely_integrable_vector_abs)
-    apply auto
-    done
-next
-  assume ?r
-  {
-    presume lem: "\<And>f::'n \<Rightarrow> 'm. f integrable_on UNIV \<Longrightarrow>
-      (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV \<Longrightarrow>
-        f absolutely_integrable_on UNIV"
-    have *: "\<And>x. (\<Sum>i\<in>Basis. \<bar>(if x \<in> s then f x else 0) \<bullet> i\<bar> *\<^sub>R i) =
-      (if x \<in> s then (\<Sum>i\<in>Basis. \<bar>f x \<bullet> i\<bar> *\<^sub>R i) else (0::'m))"
-      unfolding euclidean_eq_iff[where 'a='m]
-      by auto
-    show ?l
-      apply (subst absolutely_integrable_restrict_univ[symmetric])
-      apply (rule lem)
-      unfolding integrable_restrict_univ *
-      using \<open>?r\<close>
-      apply auto
-      done
-  }
-  fix f :: "'n \<Rightarrow> 'm"
-  assume assms: "f integrable_on UNIV" "(\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) integrable_on UNIV"
-  let ?B = "\<Sum>i\<in>Basis. integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
-  show "f absolutely_integrable_on UNIV"
-    apply (rule bounded_variation_absolutely_integrable[OF assms(1), where B="?B"])
-    apply safe
-  proof goal_cases
-    case d: (1 d)
-    note d'=division_ofD[OF d]
-    have "(\<Sum>k\<in>d. norm (integral k f)) \<le>
-      (\<Sum>k\<in>d. setsum (op \<bullet> (integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis)"
-      apply (rule setsum_mono)
-      apply (rule order_trans[OF norm_le_l1])
-      apply (rule setsum_mono)
-      unfolding lessThan_iff
-    proof -
-      fix k
-      fix i :: 'm
-      assume "k \<in> d" and i: "i \<in> Basis"
-      from d'(4)[OF this(1)] guess a b by (elim exE) note ab=this
-      show "\<bar>integral k f \<bullet> i\<bar> \<le> integral k (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m) \<bullet> i"
-        apply (rule abs_leI)
-        unfolding inner_minus_left[symmetric]
-        defer
-        apply (subst integral_neg[symmetric])
-        apply (rule_tac[1-2] integral_component_le[OF i])
-        using integrable_on_subcbox[OF assms(1),of a b]
-          integrable_on_subcbox[OF assms(2),of a b] i  integrable_neg
-        unfolding ab
-        apply auto
-        done
-    qed
-    also have "\<dots> \<le> setsum (op \<bullet> (integral UNIV (\<lambda>x. (\<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i)::'m))) Basis"
-      apply (subst setsum.commute)
-      apply (rule setsum_mono)
-    proof goal_cases
-      case (1 j)
-      have *: "(\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) integrable_on \<Union>d"
-        using integrable_on_subdivision[OF d assms(2)] by auto
-      have "(\<Sum>i\<in>d. integral i (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j) =
-        integral (\<Union>d) (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
-        unfolding inner_setsum_left[symmetric] integral_combine_division_topdown[OF * d] ..
-      also have "\<dots> \<le> integral UNIV (\<lambda>x. \<Sum>i\<in>Basis. \<bar>f x\<bullet>i\<bar> *\<^sub>R i::'m) \<bullet> j"
-        apply (rule integral_subset_component_le)
-        using assms * \<open>j \<in> Basis\<close>
-        apply auto
-        done
-      finally show ?case .
-    qed
-    finally show ?case .
-  qed
-qed
-
-lemma nonnegative_absolutely_integrable:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "\<forall>x\<in>s. \<forall>i\<in>Basis. 0 \<le> f x \<bullet> i"
-    and "f integrable_on s"
-  shows "f absolutely_integrable_on s"
-  unfolding absolutely_integrable_abs_eq
-  apply rule
-  apply (rule assms)thm integrable_eq
-  apply (rule integrable_eq[of _ f])
-  using assms
-  apply (auto simp: euclidean_eq_iff[where 'a='m])
-  done
-
-lemma absolutely_integrable_integrable_bound:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "\<forall>x\<in>s. norm (f x) \<le> g x"
-    and "f integrable_on s"
-    and "g integrable_on s"
-  shows "f absolutely_integrable_on s"
-proof -
-  {
-    presume *: "\<And>f::'n \<Rightarrow> 'm. \<And>g. \<forall>x. norm (f x) \<le> g x \<Longrightarrow> f integrable_on UNIV \<Longrightarrow>
-      g integrable_on UNIV \<Longrightarrow> f absolutely_integrable_on UNIV"
-    show ?thesis
-      apply (subst absolutely_integrable_restrict_univ[symmetric])
-      apply (rule *[of _ "\<lambda>x. if x\<in>s then g x else 0"])
-      using assms
-      unfolding integrable_restrict_univ
-      apply auto
-      done
-  }
-  fix g
-  fix f :: "'n \<Rightarrow> 'm"
-  assume assms: "\<forall>x. norm (f x) \<le> g x" "f integrable_on UNIV" "g integrable_on UNIV"
-  show "f absolutely_integrable_on UNIV"
-    apply (rule bounded_variation_absolutely_integrable[OF assms(2),where B="integral UNIV g"])
-    apply safe
-  proof goal_cases
-    case d: (1 d)
-    note d'=division_ofD[OF d]
-    have "(\<Sum>k\<in>d. norm (integral k f)) \<le> (\<Sum>k\<in>d. integral k g)"
-      apply (rule setsum_mono)
-      apply (rule integral_norm_bound_integral)
-      apply (drule_tac[!] d'(4))
-      apply safe
-      apply (rule_tac[1-2] integrable_on_subcbox)
-      using assms
-      apply auto
-      done
-    also have "\<dots> = integral (\<Union>d) g"
-      apply (rule integral_combine_division_bottomup[symmetric])
-      apply (rule d)
-      apply safe
-      apply (drule d'(4))
-      apply safe
-      apply (rule integrable_on_subcbox[OF assms(3)])
-      apply auto
-      done
-    also have "\<dots> \<le> integral UNIV g"
-      apply (rule integral_subset_le)
-      defer
-      apply (rule integrable_on_subdivision[OF d,of _ UNIV])
-      prefer 4
-      apply rule
-      apply (rule_tac y="norm (f x)" in order_trans)
-      using assms
-      apply auto
-      done
-    finally show ?case .
-  qed
-qed
-
-lemma absolutely_integrable_absolutely_integrable_bound:
-  fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-    and g :: "'n::euclidean_space \<Rightarrow> 'p::euclidean_space"
-  assumes "\<forall>x\<in>s. norm (f x) \<le> norm (g x)"
-    and "f integrable_on s"
-    and "g absolutely_integrable_on s"
-  shows "f absolutely_integrable_on s"
-  apply (rule absolutely_integrable_integrable_bound[of s f "\<lambda>x. norm (g x)"])
-  using assms
-  apply auto
-  done
-
-lemma absolutely_integrable_inf_real:
-  assumes "finite k"
-    and "k \<noteq> {}"
-    and "\<forall>i\<in>k. (\<lambda>x. (fs x i)::real) absolutely_integrable_on s"
-  shows "(\<lambda>x. (Inf ((fs x) ` k))) absolutely_integrable_on s"
-  using assms
-proof induct
-  case (insert a k)
-  let ?P = "(\<lambda>x.
-    if fs x ` k = {} then fs x a
-    else min (fs x a) (Inf (fs x ` k))) absolutely_integrable_on s"
-  show ?case
-    unfolding image_insert
-    apply (subst Inf_insert_finite)
-    apply (rule finite_imageI[OF insert(1)])
-  proof (cases "k = {}")
-    case True
-    then show ?P
-      apply (subst if_P)
-      defer
-      apply (rule insert(5)[rule_format])
-      apply auto
-      done
-  next
-    case False
-    then show ?P
-      apply (subst if_not_P)
-      defer
-      apply (rule absolutely_integrable_min[where 'n=real, unfolded Basis_real_def, simplified])
-      defer
-      apply(rule insert(3)[OF False])
-      using insert(5)
-      apply auto
-      done
-  qed
-next
-  case empty
-  then show ?case by auto
-qed
-
-lemma absolutely_integrable_sup_real:
-  assumes "finite k"
-    and "k \<noteq> {}"
-    and "\<forall>i\<in>k. (\<lambda>x. (fs x i)::real) absolutely_integrable_on s"
-  shows "(\<lambda>x. (Sup ((fs x) ` k))) absolutely_integrable_on s"
-  using assms
-proof induct
-  case (insert a k)
-  let ?P = "(\<lambda>x.
-    if fs x ` k = {} then fs x a
-    else max (fs x a) (Sup (fs x ` k))) absolutely_integrable_on s"
-  show ?case
-    unfolding image_insert
-    apply (subst Sup_insert_finite)
-    apply (rule finite_imageI[OF insert(1)])
-  proof (cases "k = {}")
-    case True
-    then show ?P
-      apply (subst if_P)
-      defer
-      apply (rule insert(5)[rule_format])
-      apply auto
-      done
-  next
-    case False
-    then show ?P
-      apply (subst if_not_P)
-      defer
-      apply (rule absolutely_integrable_max[where 'n=real, unfolded Basis_real_def, simplified])
-      defer
-      apply (rule insert(3)[OF False])
-      using insert(5)
-      apply auto
-      done
-  qed
-qed auto
-
-
-subsection \<open>differentiation under the integral sign\<close>
-
-lemma tube_lemma:
-  assumes "compact K"
-  assumes "open W"
-  assumes "{x0} \<times> K \<subseteq> W"
-  shows "\<exists>X0. x0 \<in> X0 \<and> open X0 \<and> X0 \<times> K \<subseteq> W"
-proof -
-  {
-    fix y assume "y \<in> K"
-    then have "(x0, y) \<in> W" using assms by auto
-    with \<open>open W\<close>
-    have "\<exists>X0 Y. open X0 \<and> open Y \<and> x0 \<in> X0 \<and> y \<in> Y \<and> X0 \<times> Y \<subseteq> W"
-      by (rule open_prod_elim) blast
-  }
-  then obtain X0 Y where
-    *: "\<forall>y \<in> K. open (X0 y) \<and> open (Y y) \<and> x0 \<in> X0 y \<and> y \<in> Y y \<and> X0 y \<times> Y y \<subseteq> W"
-    by metis
-  from * have "\<forall>t\<in>Y ` K. open t" "K \<subseteq> \<Union>(Y ` K)" by auto
-  with \<open>compact K\<close> obtain CC where CC: "CC \<subseteq> Y ` K" "finite CC" "K \<subseteq> \<Union>CC"
-    by (rule compactE)
-  then obtain c where c: "\<And>C. C \<in> CC \<Longrightarrow> c C \<in> K \<and> C = Y (c C)"
-    by (force intro!: choice)
-  with * CC show ?thesis
-    by (force intro!: exI[where x="\<Inter>C\<in>CC. X0 (c C)"]) (* SLOW *)
-qed
-
-lemma continuous_on_prod_compactE:
-  fixes fx::"'a::topological_space \<times> 'b::topological_space \<Rightarrow> 'c::metric_space"
-    and e::real
-  assumes cont_fx: "continuous_on (U \<times> C) fx"
-  assumes "compact C"
-  assumes [intro]: "x0 \<in> U"
-  notes [continuous_intros] = continuous_on_compose2[OF cont_fx]
-  assumes "e > 0"
-  obtains X0 where "x0 \<in> X0" "open X0"
-    "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
-proof -
-  define psi where "psi = (\<lambda>(x, t). dist (fx (x, t)) (fx (x0, t)))"
-  define W0 where "W0 = {(x, t) \<in> U \<times> C. psi (x, t) < e}"
-  have W0_eq: "W0 = psi -` {..<e} \<inter> U \<times> C"
-    by (auto simp: vimage_def W0_def)
-  have "open {..<e}" by simp
-  have "continuous_on (U \<times> C) psi"
-    by (auto intro!: continuous_intros simp: psi_def split_beta')
-  from this[unfolded continuous_on_open_invariant, rule_format, OF \<open>open {..<e}\<close>]
-  obtain W where W: "open W" "W \<inter> U \<times> C = W0 \<inter> U \<times> C"
-    unfolding W0_eq by blast
-  have "{x0} \<times> C \<subseteq> W \<inter> U \<times> C"
-    unfolding W
-    by (auto simp: W0_def psi_def \<open>0 < e\<close>)
-  then have "{x0} \<times> C \<subseteq> W" by blast
-  from tube_lemma[OF \<open>compact C\<close> \<open>open W\<close> this]
-  obtain X0 where X0: "x0 \<in> X0" "open X0" "X0 \<times> C \<subseteq> W"
-    by blast
-
-  have "\<forall>x\<in>X0 \<inter> U. \<forall>t \<in> C. dist (fx (x, t)) (fx (x0, t)) \<le> e"
-  proof safe
-    fix x assume x: "x \<in> X0" "x \<in> U"
-    fix t assume t: "t \<in> C"
-    have "dist (fx (x, t)) (fx (x0, t)) = psi (x, t)"
-      by (auto simp: psi_def)
-    also
-    {
-      have "(x, t) \<in> X0 \<times> C"
-        using t x
-        by auto
-      also note \<open>\<dots> \<subseteq> W\<close>
-      finally have "(x, t) \<in> W" .
-      with t x have "(x, t) \<in> W \<inter> U \<times> C"
-        by blast
-      also note \<open>W \<inter> U \<times> C = W0 \<inter> U \<times> C\<close>
-      finally  have "psi (x, t) < e"
-        by (auto simp: W0_def)
-    }
-    finally show "dist (fx (x, t)) (fx (x0, t)) \<le> e" by simp
-  qed
-  from X0(1,2) this show ?thesis ..
-qed
-
-lemma integral_continuous_on_param:
-  fixes f::"'a::topological_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
-  assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). f x t)"
-  shows "continuous_on U (\<lambda>x. integral (cbox a b) (f x))"
-proof cases
-  assume "content (cbox a b) \<noteq> 0"
-  then have ne: "cbox a b \<noteq> {}" by auto
-
-  note [continuous_intros] =
-    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
-      unfolded split_beta fst_conv snd_conv]
-
-  show ?thesis
-    unfolding continuous_on_def
-  proof (safe intro!: tendstoI)
-    fix e'::real and x
-    assume "e' > 0"
-    define e where "e = e' / (content (cbox a b) + 1)"
-    have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
-    assume "x \<in> U"
-    from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x \<in> U\<close> \<open>0 < e\<close>]
-    obtain X0 where X0: "x \<in> X0" "open X0"
-      and fx_bound: "\<And>y t. y \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (f y t - f x t) \<le> e"
-      unfolding split_beta fst_conv snd_conv dist_norm
-      by metis
-    have "\<forall>\<^sub>F y in at x within U. y \<in> X0 \<inter> U"
-      using X0(1) X0(2) eventually_at_topological by auto
-    then show "\<forall>\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'"
-    proof eventually_elim
-      case (elim y)
-      have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) =
-        norm (integral (cbox a b) (\<lambda>t. f y t - f x t))"
-        using elim \<open>x \<in> U\<close>
-        unfolding dist_norm
-        by (subst integral_diff)
-           (auto intro!: integrable_continuous continuous_intros)
-      also have "\<dots> \<le> e * content (cbox a b)"
-        using elim \<open>x \<in> U\<close>
-        by (intro integrable_bound)
-           (auto intro!: fx_bound \<open>x \<in> U \<close> less_imp_le[OF \<open>0 < e\<close>]
-              integrable_continuous continuous_intros)
-      also have "\<dots> < e'"
-        using \<open>0 < e'\<close> \<open>e > 0\<close>
-        by (auto simp: e_def divide_simps)
-      finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" .
-    qed
-  qed
-qed (auto intro!: continuous_on_const)
-
-lemma eventually_closed_segment:
-  fixes x0::"'a::real_normed_vector"
-  assumes "open X0" "x0 \<in> X0"
-  shows "\<forall>\<^sub>F x in at x0 within U. closed_segment x0 x \<subseteq> X0"
-proof -
-  from openE[OF assms]
-  obtain e where e: "0 < e" "ball x0 e \<subseteq> X0" .
-  then have "\<forall>\<^sub>F x in at x0 within U. x \<in> ball x0 e"
-    by (auto simp: dist_commute eventually_at)
-  then show ?thesis
-  proof eventually_elim
-    case (elim x)
-    have "x0 \<in> ball x0 e" using \<open>e > 0\<close> by simp
-    from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim]
-    have "closed_segment x0 x \<subseteq> ball x0 e" .
-    also note \<open>\<dots> \<subseteq> X0\<close>
-    finally show ?case .
-  qed
-qed
-
-lemma leibniz_rule:
-  fixes f::"'a::banach \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
-  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
-    ((\<lambda>x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)"
-  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> f x integrable_on cbox a b"
-  assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
-  assumes [intro]: "x0 \<in> U"
-  assumes "convex U"
-  shows
-    "((\<lambda>x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
-    (is "(?F has_derivative ?dF) _")
-proof cases
-  assume "content (cbox a b) \<noteq> 0"
-  then have ne: "cbox a b \<noteq> {}" by auto
-  note [continuous_intros] =
-    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
-      unfolded split_beta fst_conv snd_conv]
-  show ?thesis
-  proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist)
-    have cont_f1: "\<And>t. t \<in> cbox a b \<Longrightarrow> continuous_on U (\<lambda>x. f x t)"
-      by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx)
-    note [continuous_intros] = continuous_on_compose2[OF cont_f1]
-    fix e'::real
-    assume "e' > 0"
-    define e where "e = e' / (content (cbox a b) + 1)"
-    have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos)
-    from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x0 \<in> U\<close> \<open>e > 0\<close>]
-    obtain X0 where X0: "x0 \<in> X0" "open X0"
-      and fx_bound: "\<And>x t. x \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (fx x t - fx x0 t) \<le> e"
-      unfolding split_beta fst_conv snd_conv
-      by (metis dist_norm)
-
-    note eventually_closed_segment[OF \<open>open X0\<close> \<open>x0 \<in> X0\<close>, of U]
-    moreover
-    have "\<forall>\<^sub>F x in at x0 within U. x \<in> X0"
-      using \<open>open X0\<close> \<open>x0 \<in> X0\<close> eventually_at_topological by blast
-    moreover have "\<forall>\<^sub>F x in at x0 within U. x \<noteq> x0"
-      by (auto simp: eventually_at_filter)
-    moreover have "\<forall>\<^sub>F x in at x0 within U. x \<in> U"
-      by (auto simp: eventually_at_filter)
-    ultimately
-    show "\<forall>\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'"
-    proof eventually_elim
-      case (elim x)
-      from elim have "0 < norm (x - x0)" by simp
-      have "closed_segment x0 x \<subseteq> U"
-        by (rule \<open>convex U\<close>[unfolded convex_contains_segment, rule_format, OF \<open>x0 \<in> U\<close> \<open>x \<in> U\<close>])
-      from elim have [intro]: "x \<in> U" by auto
-
-      have "?F x - ?F x0 - ?dF (x - x0) =
-        integral (cbox a b) (\<lambda>y. f x y - f x0 y - fx x0 y (x - x0))"
-        (is "_ = ?id")
-        using \<open>x \<noteq> x0\<close>
-        by (subst blinfun_apply_integral integral_diff,
-            auto intro!: integrable_diff integrable_f2 continuous_intros
-              intro: integrable_continuous)+
-      also
-      {
-        fix t assume t: "t \<in> (cbox a b)"
-        have seg: "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R (x - x0) \<in> X0 \<inter> U"
-          using \<open>closed_segment x0 x \<subseteq> U\<close>
-            \<open>closed_segment x0 x \<subseteq> X0\<close>
-          by (force simp: closed_segment_def algebra_simps)
-        from t have deriv:
-          "((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)"
-          if "y \<in> X0 \<inter> U" for y
-          unfolding has_vector_derivative_def[symmetric]
-          using that \<open>x \<in> X0\<close>
-          by (intro has_derivative_within_subset[OF fx]) auto
-        have "\<forall>x \<in> X0 \<inter> U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \<le> e"
-          using fx_bound t
-          by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric])
-        from differentiable_bound_linearization[OF seg deriv this] X0
-        have "norm (f x t - f x0 t - fx x0 t (x - x0)) \<le> e * norm (x - x0)"
-          by (auto simp add: ac_simps)
-      }
-      then have "norm ?id \<le> integral (cbox a b) (\<lambda>_. e * norm (x - x0))"
-        by (intro integral_norm_bound_integral)
-          (auto intro!: continuous_intros integrable_diff integrable_f2
-            intro: integrable_continuous)
-      also have "\<dots> = content (cbox a b) * e * norm (x - x0)"
-        by simp
-      also have "\<dots> < e' * norm (x - x0)"
-        using \<open>e' > 0\<close> content_pos_le[of a b]
-        by (intro mult_strict_right_mono[OF _ \<open>0 < norm (x - x0)\<close>])
-          (auto simp: divide_simps e_def)
-      finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" .
-      then show ?case
-        by (auto simp: divide_simps)
-    qed
-  qed (rule blinfun.bounded_linear_right)
-qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps)
-
-lemma
-  has_vector_derivative_eq_has_derivative_blinfun:
-  "(f has_vector_derivative f') (at x within U) \<longleftrightarrow>
-    (f has_derivative blinfun_scaleR_left f') (at x within U)"
-  by (simp add: has_vector_derivative_def)
-
-lemma leibniz_rule_vector_derivative:
-  fixes f::"real \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
-  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow>
-      ((\<lambda>x. f x t) has_vector_derivative (fx x t)) (at x within U)"
-  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
-  assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). fx x t)"
-  assumes U: "x0 \<in> U" "convex U"
-  shows "((\<lambda>x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0))
-      (at x0 within U)"
-proof -
-  note [continuous_intros] =
-    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
-      unfolded split_beta fst_conv snd_conv]
-  have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) =
-    integral (cbox a b) (\<lambda>t. blinfun_scaleR_left (fx x0 t))"
-    by (subst integral_linear[symmetric])
-       (auto simp: has_vector_derivative_def o_def
-         intro!: integrable_continuous U continuous_intros bounded_linear_intros)
-  show ?thesis
-    unfolding has_vector_derivative_eq_has_derivative_blinfun
-    apply (rule has_derivative_eq_rhs)
-    apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_scaleR_left (fx x t)"])
-    using fx cont_fx
-    apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros)
-    done
-qed
-
-lemma
-  has_field_derivative_eq_has_derivative_blinfun:
-  "(f has_field_derivative f') (at x within U) \<longleftrightarrow> (f has_derivative blinfun_mult_right f') (at x within U)"
-  by (simp add: has_field_derivative_def)
-
-lemma leibniz_rule_field_derivative:
-  fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
-  assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)"
-  assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b"
-  assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)"
-  assumes U: "x0 \<in> U" "convex U"
-  shows "((\<lambda>x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)"
-proof -
-  note [continuous_intros] =
-    continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x,
-      unfolded split_beta fst_conv snd_conv]
-  have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) =
-    integral (cbox a b) (\<lambda>t. blinfun_mult_right (fx x0 t))"
-    by (subst integral_linear[symmetric])
-      (auto simp: has_vector_derivative_def o_def
-        intro!: integrable_continuous U continuous_intros bounded_linear_intros)
-  show ?thesis
-    unfolding has_field_derivative_eq_has_derivative_blinfun
-    apply (rule has_derivative_eq_rhs)
-    apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_mult_right (fx x t)"])
-    using fx cont_fx
-    apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros)
-    done
-qed
-
-
-subsection \<open>Exchange uniform limit and integral\<close>
-
-lemma
-  uniform_limit_integral:
-  fixes f::"'a \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach"
-  assumes u: "uniform_limit (cbox a b) f g F"
-  assumes c: "\<And>n. continuous_on (cbox a b) (f n)"
-  assumes [simp]: "F \<noteq> bot"
-  obtains I J where
-    "\<And>n. (f n has_integral I n) (cbox a b)"
-    "(g has_integral J) (cbox a b)"
-    "(I \<longlongrightarrow> J) F"
-proof -
-  have fi[simp]: "f n integrable_on (cbox a b)" for n
-    by (auto intro!: integrable_continuous assms)
-  then obtain I where I: "\<And>n. (f n has_integral I n) (cbox a b)"
-    by atomize_elim (auto simp: integrable_on_def intro!: choice)
-
-  moreover
-
-  have gi[simp]: "g integrable_on (cbox a b)"
-    by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c)
-  then obtain J where J: "(g has_integral J) (cbox a b)"
-    by blast
-
-  moreover
-
-  have "(I \<longlongrightarrow> J) F"
-  proof cases
-    assume "content (cbox a b) = 0"
-    hence "I = (\<lambda>_. 0)" "J = 0"
-      by (auto intro!: has_integral_unique I J)
-    thus ?thesis by simp
-  next
-    assume content_nonzero: "content (cbox a b) \<noteq> 0"
-    show ?thesis
-    proof (rule tendstoI)
-      fix e::real
-      assume "e > 0"
-      define e' where "e' = e / 2"
-      with \<open>e > 0\<close> have "e' > 0" by simp
-      then have "\<forall>\<^sub>F n in F. \<forall>x\<in>cbox a b. norm (f n x - g x) < e' / content (cbox a b)"
-        using u content_nonzero content_pos_le[of a b]
-        by (auto simp: uniform_limit_iff dist_norm)
-      then show "\<forall>\<^sub>F n in F. dist (I n) J < e"
-      proof eventually_elim
-        case (elim n)
-        have "I n = integral (cbox a b) (f n)"
-            "J = integral (cbox a b) g"
-          using I[of n] J by (simp_all add: integral_unique)
-        then have "dist (I n) J = norm (integral (cbox a b) (\<lambda>x. f n x - g x))"
-          by (simp add: integral_diff dist_norm)
-        also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. (e' / content (cbox a b)))"
-          using elim
-          by (intro integral_norm_bound_integral)
-            (auto intro!: integrable_diff absolutely_integrable_onI)
-        also have "\<dots> < e"
-          using \<open>0 < e\<close>
-          by (simp add: e'_def)
-        finally show ?case .
-      qed
-    qed
-  qed
-  ultimately show ?thesis ..
-qed
-
-
-subsection \<open>Dominated convergence\<close>
-
-context
-begin
-
-private lemma dominated_convergence_real:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real"
-  assumes "\<And>k. (f k) integrable_on s" "h integrable_on s"
-    and "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
-    and "\<forall>x \<in> s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-  shows "g integrable_on s \<and> ((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-proof
-  have bdd_below[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_below {f n x |n. P n}"
-  proof (safe intro!: bdd_belowI)
-    fix n x show "x \<in> s \<Longrightarrow> - h x \<le> f n x"
-      using assms(3)[rule_format, of x n] by simp
-  qed
-  have bdd_above[simp]: "\<And>x P. x \<in> s \<Longrightarrow> bdd_above {f n x |n. P n}"
-  proof (safe intro!: bdd_aboveI)
-    fix n x show "x \<in> s \<Longrightarrow> f n x \<le> h x"
-      using assms(3)[rule_format, of x n] by simp
-  qed
-
-  have "\<And>m. (\<lambda>x. Inf {f j x |j. m \<le> j}) integrable_on s \<and>
-    ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) \<longlongrightarrow>
-    integral s (\<lambda>x. Inf {f j x |j. m \<le> j}))sequentially"
-  proof (rule monotone_convergence_decreasing, safe)
-    fix m :: nat
-    show "bounded {integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}}) |k. True}"
-      unfolding bounded_iff
-      apply (rule_tac x="integral s h" in exI)
-    proof safe
-      fix k :: nat
-      show "norm (integral s (\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}})) \<le> integral s h"
-        apply (rule integral_norm_bound_integral)
-        unfolding simple_image
-        apply (rule absolutely_integrable_onD)
-        apply (rule absolutely_integrable_inf_real)
-        prefer 5
-        unfolding real_norm_def
-        apply rule
-        apply (rule cInf_abs_ge)
-        prefer 5
-        apply rule
-        apply (rule_tac g = h in absolutely_integrable_integrable_bound)
-        using assms
-        unfolding real_norm_def
-        apply auto
-        done
-    qed
-    fix k :: nat
-    show "(\<lambda>x. Inf {f j x |j. j \<in> {m..m + k}}) integrable_on s"
-      unfolding simple_image
-      apply (rule absolutely_integrable_onD)
-      apply (rule absolutely_integrable_inf_real)
-      prefer 3
-      using absolutely_integrable_integrable_bound[OF assms(3,1,2)]
-      apply auto
-      done
-    fix x
-    assume x: "x \<in> s"
-    show "Inf {f j x |j. j \<in> {m..m + Suc k}} \<le> Inf {f j x |j. j \<in> {m..m + k}}"
-      by (rule cInf_superset_mono) auto
-    let ?S = "{f j x| j. m \<le> j}"
-    show "((\<lambda>k. Inf {f j x |j. j \<in> {m..m + k}}) \<longlongrightarrow> Inf ?S) sequentially"
-    proof (rule LIMSEQ_I, goal_cases)
-      case r: (1 r)
-
-      have "\<exists>y\<in>?S. y < Inf ?S + r"
-        by (subst cInf_less_iff[symmetric]) (auto simp: \<open>x\<in>s\<close> r)
-      then obtain N where N: "f N x < Inf ?S + r" "m \<le> N"
-        by blast
-
-      show ?case
-        apply (rule_tac x=N in exI)
-        apply safe
-      proof goal_cases
-        case prems: (1 n)
-        have *: "\<And>y ix. y < Inf ?S + r \<longrightarrow> Inf ?S \<le> ix \<longrightarrow> ix \<le> y \<longrightarrow> \<bar>ix - Inf ?S\<bar> < r"
-          by arith
-        show ?case
-          unfolding real_norm_def
-            apply (rule *[rule_format, OF N(1)])
-            apply (rule cInf_superset_mono, auto simp: \<open>x\<in>s\<close>) []
-            apply (rule cInf_lower)
-            using N prems
-            apply auto []
-            apply simp
-            done
-      qed
-    qed
-  qed
-  note dec1 = conjunctD2[OF this]
-
-  have "\<And>m. (\<lambda>x. Sup {f j x |j. m \<le> j}) integrable_on s \<and>
-    ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}})) \<longlongrightarrow>
-    integral s (\<lambda>x. Sup {f j x |j. m \<le> j})) sequentially"
-  proof (rule monotone_convergence_increasing,safe)
-    fix m :: nat
-    show "bounded {integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}}) |k. True}"
-      unfolding bounded_iff
-      apply (rule_tac x="integral s h" in exI)
-    proof safe
-      fix k :: nat
-      show "norm (integral s (\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}})) \<le> integral s h"
-        apply (rule integral_norm_bound_integral) unfolding simple_image
-        apply (rule absolutely_integrable_onD)
-        apply(rule absolutely_integrable_sup_real)
-        prefer 5 unfolding real_norm_def
-        apply rule
-        apply (rule cSup_abs_le)
-        using assms
-        apply (force simp add: )
-        prefer 4
-        apply rule
-        apply (rule_tac g=h in absolutely_integrable_integrable_bound)
-        using assms
-        unfolding real_norm_def
-        apply auto
-        done
-    qed
-    fix k :: nat
-    show "(\<lambda>x. Sup {f j x |j. j \<in> {m..m + k}}) integrable_on s"
-      unfolding simple_image
-      apply (rule absolutely_integrable_onD)
-      apply (rule absolutely_integrable_sup_real)
-      prefer 3
-      using absolutely_integrable_integrable_bound[OF assms(3,1,2)]
-      apply auto
-      done
-    fix x
-    assume x: "x\<in>s"
-    show "Sup {f j x |j. j \<in> {m..m + Suc k}} \<ge> Sup {f j x |j. j \<in> {m..m + k}}"
-      by (rule cSup_subset_mono) auto
-    let ?S = "{f j x| j. m \<le> j}"
-    show "((\<lambda>k. Sup {f j x |j. j \<in> {m..m + k}}) \<longlongrightarrow> Sup ?S) sequentially"
-    proof (rule LIMSEQ_I, goal_cases)
-      case r: (1 r)
-      have "\<exists>y\<in>?S. Sup ?S - r < y"
-        by (subst less_cSup_iff[symmetric]) (auto simp: r \<open>x\<in>s\<close>)
-      then obtain N where N: "Sup ?S - r < f N x" "m \<le> N"
-        by blast
-
-      show ?case
-        apply (rule_tac x=N in exI)
-        apply safe
-      proof goal_cases
-        case prems: (1 n)
-        have *: "\<And>y ix. Sup ?S - r < y \<longrightarrow> ix \<le> Sup ?S \<longrightarrow> y \<le> ix \<longrightarrow> \<bar>ix - Sup ?S\<bar> < r"
-          by arith
-        show ?case
-          apply simp
-          apply (rule *[rule_format, OF N(1)])
-          apply (rule cSup_subset_mono, auto simp: \<open>x\<in>s\<close>) []
-          apply (subst cSup_upper)
-          using N prems
-          apply auto
-          done
-      qed
-    qed
-  qed
-  note inc1 = conjunctD2[OF this]
-
-  have "g integrable_on s \<and>
-    ((\<lambda>k. integral s (\<lambda>x. Inf {f j x |j. k \<le> j})) \<longlongrightarrow> integral s g) sequentially"
-    apply (rule monotone_convergence_increasing,safe)
-    apply fact
-  proof -
-    show "bounded {integral s (\<lambda>x. Inf {f j x |j. k \<le> j}) |k. True}"
-      unfolding bounded_iff apply(rule_tac x="integral s h" in exI)
-    proof safe
-      fix k :: nat
-      show "norm (integral s (\<lambda>x. Inf {f j x |j. k \<le> j})) \<le> integral s h"
-        apply (rule integral_norm_bound_integral)
-        apply fact+
-        unfolding real_norm_def
-        apply rule
-        apply (rule cInf_abs_ge)
-        using assms(3)
-        apply auto
-        done
-    qed
-    fix k :: nat and x
-    assume x: "x \<in> s"
-
-    have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
-    show "Inf {f j x |j. k \<le> j} \<le> Inf {f j x |j. Suc k \<le> j}"
-      by (intro cInf_superset_mono) (auto simp: \<open>x\<in>s\<close>)
-
-    show "(\<lambda>k::nat. Inf {f j x |j. k \<le> j}) \<longlonglongrightarrow> g x"
-    proof (rule LIMSEQ_I, goal_cases)
-      case r: (1 r)
-      then have "0<r/2"
-        by auto
-      from assms(4)[THEN bspec, THEN LIMSEQ_D, OF x this] guess N .. note N = this
-      show ?case
-        apply (rule_tac x=N in exI)
-        apply safe
-        unfolding real_norm_def
-        apply (rule le_less_trans[of _ "r/2"])
-        apply (rule cInf_asclose)
-        apply safe
-        defer
-        apply (rule less_imp_le)
-        using N r
-        apply auto
-        done
-    qed
-  qed
-  note inc2 = conjunctD2[OF this]
-
-  have "g integrable_on s \<and>
-    ((\<lambda>k. integral s (\<lambda>x. Sup {f j x |j. k \<le> j})) \<longlongrightarrow> integral s g) sequentially"
-    apply (rule monotone_convergence_decreasing,safe)
-    apply fact
-  proof -
-    show "bounded {integral s (\<lambda>x. Sup {f j x |j. k \<le> j}) |k. True}"
-      unfolding bounded_iff
-      apply (rule_tac x="integral s h" in exI)
-    proof safe
-      fix k :: nat
-      show "norm (integral s (\<lambda>x. Sup {f j x |j. k \<le> j})) \<le> integral s h"
-        apply (rule integral_norm_bound_integral)
-        apply fact+
-        unfolding real_norm_def
-        apply rule
-        apply (rule cSup_abs_le)
-        using assms(3)
-        apply auto
-        done
-    qed
-    fix k :: nat
-    fix x
-    assume x: "x \<in> s"
-
-    show "Sup {f j x |j. k \<le> j} \<ge> Sup {f j x |j. Suc k \<le> j}"
-      by (rule cSup_subset_mono) (auto simp: \<open>x\<in>s\<close>)
-    show "((\<lambda>k. Sup {f j x |j. k \<le> j}) \<longlongrightarrow> g x) sequentially"
-    proof (rule LIMSEQ_I, goal_cases)
-      case r: (1 r)
-      then have "0<r/2"
-        by auto
-      from assms(4)[THEN bspec, THEN LIMSEQ_D, OF x this] guess N .. note N=this
-      show ?case
-        apply (rule_tac x=N in exI,safe)
-        unfolding real_norm_def
-        apply (rule le_less_trans[of _ "r/2"])
-        apply (rule cSup_asclose)
-        apply safe
-        defer
-        apply (rule less_imp_le)
-        using N r
-        apply auto
-        done
-    qed
-  qed
-  note dec2 = conjunctD2[OF this]
-
-  show "g integrable_on s" by fact
-  show "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-  proof (rule LIMSEQ_I, goal_cases)
-    case r: (1 r)
-    from LIMSEQ_D [OF inc2(2) r] guess N1 .. note N1=this[unfolded real_norm_def]
-    from LIMSEQ_D [OF dec2(2) r] guess N2 .. note N2=this[unfolded real_norm_def]
-    show ?case
-    proof (rule_tac x="N1+N2" in exI, safe)
-      fix n
-      assume n: "n \<ge> N1 + N2"
-      have *: "\<And>i0 i i1 g. \<bar>i0 - g\<bar> < r \<longrightarrow> \<bar>i1 - g\<bar> < r \<longrightarrow> i0 \<le> i \<longrightarrow> i \<le> i1 \<longrightarrow> \<bar>i - g\<bar> < r"
-        by arith
-      show "norm (integral s (f n) - integral s g) < r"
-        unfolding real_norm_def
-      proof (rule *[rule_format,OF N1[rule_format] N2[rule_format], of n n])
-        show "integral s (\<lambda>x. Inf {f j x |j. n \<le> j}) \<le> integral s (f n)"
-          by (rule integral_le[OF dec1(1) assms(1)]) (auto intro!: cInf_lower)
-        show "integral s (f n) \<le> integral s (\<lambda>x. Sup {f j x |j. n \<le> j})"
-          by (rule integral_le[OF assms(1) inc1(1)]) (auto intro!: cSup_upper)
-      qed (insert n, auto)
-    qed
-  qed
-qed
-
-lemma dominated_convergence:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes f: "\<And>k. (f k) integrable_on s" and h: "h integrable_on s"
-    and le: "\<And>k. \<forall>x \<in> s. norm (f k x) \<le> h x"
-    and conv: "\<forall>x \<in> s. ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially"
-  shows "g integrable_on s"
-    and "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-proof -
-  {
-    fix b :: 'm assume b: "b \<in> Basis"
-    have A: "(\<lambda>x. g x \<bullet> b) integrable_on s \<and>
-               (\<lambda>k. integral s (\<lambda>x. f k x \<bullet> b)) \<longlonglongrightarrow> integral s (\<lambda>x. g x \<bullet> b)"
-    proof (rule dominated_convergence_real)
-      fix k :: nat
-      from f show "(\<lambda>x. f k x \<bullet> b) integrable_on s" by (rule integrable_component)
-    next
-      from h show "h integrable_on s" .
-    next
-      fix k :: nat
-      show "\<forall>x\<in>s. norm (f k x \<bullet> b) \<le> h x"
-      proof
-        fix x assume x: "x \<in> s"
-        have "norm (f k x \<bullet> b) \<le> norm (f k x)" by (simp add: Basis_le_norm b)
-        also have "\<dots> \<le> h x" using le[of k] x by simp
-        finally show "norm (f k x \<bullet> b) \<le> h x" .
-      qed
-    next
-      from conv show "\<forall>x\<in>s. (\<lambda>k. f k x \<bullet> b) \<longlonglongrightarrow> g x \<bullet> b"
-        by (auto intro!: tendsto_inner)
-    qed
-    have B: "integral s ((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f k x \<bullet> b)) = integral s (\<lambda>x. f k x \<bullet> b) *\<^sub>R b"
-      for k by (rule integral_linear)
-               (simp_all add: f integrable_component bounded_linear_scaleR_left)
-    have C: "integral s ((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. g x \<bullet> b)) = integral s (\<lambda>x. g x \<bullet> b) *\<^sub>R b"
-      using A by (intro integral_linear)
-                 (simp_all add: integrable_component bounded_linear_scaleR_left)
-    note A B C
-  } note A = this
-
-  show "g integrable_on s" by (rule integrable_componentwise) (insert A, blast)
-  with A f show "((\<lambda>k. integral s (f k)) \<longlongrightarrow> integral s g) sequentially"
-    by (subst (1 2) integral_componentwise)
-       (auto intro!: tendsto_setsum tendsto_scaleR simp: o_def)
-qed
-
-lemma has_integral_dominated_convergence:
-  fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> 'm::euclidean_space"
-  assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s"
-    "\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) \<longlonglongrightarrow> g x"
-    and x: "y \<longlonglongrightarrow> x"
-  shows "(g has_integral x) s"
-proof -
-  have int_f: "\<And>k. (f k) integrable_on s"
-    using assms by (auto simp: integrable_on_def)
-  have "(g has_integral (integral s g)) s"
-    by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+
-  moreover have "integral s g = x"
-  proof (rule LIMSEQ_unique)
-    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> x"
-      using integral_unique[OF assms(1)] x by simp
-    show "(\<lambda>i. integral s (f i)) \<longlonglongrightarrow> integral s g"
-      by (intro dominated_convergence[OF int_f assms(2)]) fact+
-  qed
-  ultimately show ?thesis
-    by simp
-qed
-
-end
-
-
-subsection \<open>Integration by parts\<close>
-
-lemma integration_by_parts_interior_strong:
-  assumes bilinear: "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)"
-  assumes s: "finite s" and le: "a \<le> b"
-  assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-                 "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes int: "((\<lambda>x. prod (f x) (g' x)) has_integral
-                  (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
-  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
-proof -
-  interpret bounded_bilinear prod by fact
-  have "((\<lambda>x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral
-          (prod (f b) (g b) - prod (f a) (g a))) {a..b}"
-    using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le])
-                   (auto intro!: continuous_intros continuous_on has_vector_derivative)
-  from has_integral_sub[OF this int] show ?thesis by (simp add: algebra_simps)
-qed
-
-lemma integration_by_parts_interior:
-  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
-          "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-          "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
-  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
-  by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
-
-lemma integration_by_parts:
-  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
-          "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-          "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
-  shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
-  by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all)
-
-lemma integrable_by_parts_interior_strong:
-  assumes bilinear: "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)"
-  assumes s: "finite s" and le: "a \<le> b"
-  assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-                 "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes int: "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
-  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
-proof -
-  from int obtain I where "((\<lambda>x. prod (f x) (g' x)) has_integral I) {a..b}"
-    unfolding integrable_on_def by blast
-  hence "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) -
-           (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp
-  from integration_by_parts_interior_strong[OF assms(1-7) this]
-    show ?thesis unfolding integrable_on_def by blast
-qed
-
-lemma integrable_by_parts_interior:
-  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
-          "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-          "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
-  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
-  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
-
-lemma integrable_by_parts:
-  assumes "bounded_bilinear (prod :: _ \<Rightarrow> _ \<Rightarrow> 'b :: banach)" "a \<le> b"
-          "continuous_on {a..b} f" "continuous_on {a..b} g"
-  assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
-          "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-  assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
-  shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
-  by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
-
-
-subsection \<open>Integration by substitution\<close>
-
-
-lemma has_integral_substitution_strong:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
-  assumes s: "finite s" and le: "a \<le> b" "c \<le> d" "g a \<le> g b"
-      and subset: "g ` {a..b} \<subseteq> {c..d}"
-      and f [continuous_intros]: "continuous_on {c..d} f"
-      and g [continuous_intros]: "continuous_on {a..b} g"
-      and deriv [derivative_intros]:
-              "\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
-    shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
-proof -
-  let ?F = "\<lambda>x. integral {c..g x} f"
-  have cont_int: "continuous_on {a..b} ?F"
-    by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous
-          f integrable_continuous_real)+
-  have deriv: "(((\<lambda>x. integral {c..x} f) \<circ> g) has_vector_derivative g' x *\<^sub>R f (g x))
-                 (at x within {a..b})" if "x \<in> {a..b} - s" for x
-    apply (rule has_vector_derivative_eq_rhs)
-    apply (rule vector_diff_chain_within)
-    apply (subst has_field_derivative_iff_has_vector_derivative [symmetric])
-    apply (rule deriv that)+
-    apply (rule has_vector_derivative_within_subset)
-    apply (rule integral_has_vector_derivative f)+
-    using that le subset
-    apply blast+
-    done
-  have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x))
-                  (at x)" if "x \<in> {a..b} - (s \<union> {a,b})" for x
-    using deriv[of x] that by (simp add: at_within_closed_interval o_def)
-
-
-  have "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}"
-    using le cont_int s deriv cont_int
-    by (intro fundamental_theorem_of_calculus_interior_strong[of "s \<union> {a,b}"]) simp_all
-  also from subset have "g x \<in> {c..d}" if "x \<in> {a..b}" for x using that by blast
-  from this[of a] this[of b] le have "c \<le> g a" "g b \<le> d" by auto
-  hence "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f"
-    by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all
-  hence "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f"
-    by (simp add: algebra_simps)
-  finally show ?thesis .
-qed
-
-lemma has_integral_substitution:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real"
-  assumes "a \<le> b" "c \<le> d" "g a \<le> g b" "g ` {a..b} \<subseteq> {c..d}"
-      and "continuous_on {c..d} f"
-      and "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
-    shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
-  by (intro has_integral_substitution_strong[of "{}" a b c d] assms)
-     (auto intro: DERIV_continuous_on assms)
-
-
-subsection \<open>Compute a double integral using iterated integrals and switching the order of integration\<close>
-
-lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
-  by auto
-
-lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
-  by auto
-
-lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
-  by blast
-
-lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
-  by blast
-
-lemma continuous_on_imp_integrable_on_Pair1:
-  fixes f :: "_ \<Rightarrow> 'b::banach"
-  assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \<in> cbox a b"
-  shows "(\<lambda>y. f (x, y)) integrable_on (cbox c d)"
-proof -
-  have "f \<circ> (\<lambda>y. (x, y)) integrable_on (cbox c d)"
-    apply (rule integrable_continuous)
-    apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]])
-    using x
-    apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con)
-    done
-  then show ?thesis
-    by (simp add: o_def)
-qed
-
-lemma integral_integrable_2dim:
-  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
-  assumes "continuous_on (cbox (a,c) (b,d)) f"
-    shows "(\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y))) integrable_on cbox a b"
-proof (cases "content(cbox c d) = 0")
-case True
-  then show ?thesis
-    by (simp add: True integrable_const)
-next
-  case False
-  have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f"
-    by (simp add: assms compact_cbox compact_uniformly_continuous)
-  { fix x::'a and e::real
-    assume x: "x \<in> cbox a b" and e: "0 < e"
-    then have e2_gt: "0 < e / 2 / content (cbox c d)" and e2_less: "e / 2 / content (cbox c d) * content (cbox c d) < e"
-      by (auto simp: False content_lt_nz e)
-    then obtain dd
-    where dd: "\<And>x x'. \<lbrakk>x\<in>cbox (a, c) (b, d); x'\<in>cbox (a, c) (b, d); norm (x' - x) < dd\<rbrakk>
-                       \<Longrightarrow> norm (f x' - f x) \<le> e / (2 * content (cbox c d))"  "dd>0"
-      using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e / (2 * content (cbox c d))"]
-      by (auto simp: dist_norm intro: less_imp_le)
-    have "\<exists>delta>0. \<forall>x'\<in>cbox a b. norm (x' - x) < delta \<longrightarrow> norm (integral (cbox c d) (\<lambda>u. f (x', u) - f (x, u))) < e"
-      apply (rule_tac x=dd in exI)
-      using dd e2_gt assms x
-      apply clarify
-      apply (rule le_less_trans [OF _ e2_less])
-      apply (rule integrable_bound)
-      apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1)
-      done
-  } note * = this
-  show ?thesis
-    apply (rule integrable_continuous)
-    apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms])
-    done
-qed
-
-lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk>
-            \<Longrightarrow> norm(y - x) \<le> e"
-using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"]
-  by (simp add: add_diff_add)
-
-lemma integral_split:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
-  assumes f: "f integrable_on (cbox a b)"
-      and k: "k \<in> Basis"
-  shows "integral (cbox a b) f =
-           integral (cbox a b \<inter> {x. x\<bullet>k \<le> c}) f +
-           integral (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) f"
-  apply (rule integral_unique [OF has_integral_split [where c=c]])
-  using k f
-  apply (auto simp: has_integral_integral [symmetric])
-  done
-
-lemma integral_swap_operative:
-  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
-  assumes f: "continuous_on s f" and e: "0 < e"
-    shows "comm_monoid.operative (op \<and>) True
-           (\<lambda>k. \<forall>a b c d.
-                cbox (a,c) (b,d) \<subseteq> k \<and> cbox (a,c) (b,d) \<subseteq> s
-                \<longrightarrow> norm(integral (cbox (a,c) (b,d)) f -
-                         integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f((x,y)))))
-                    \<le> e * content (cbox (a,c) (b,d)))"
-proof (auto simp: comm_monoid.operative_def[OF comm_monoid_and])
-  fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b
-  assume c0: "content (cbox (a, c) (b, d)) = 0"
-     and cb1: "cbox (u, w) (v, z) \<subseteq> cbox (a, c) (b, d)"
-     and cb2: "cbox (u, w) (v, z) \<subseteq> s"
-  have c0': "content (cbox (u, w) (v, z)) = 0"
-    by (fact content_0_subset [OF c0 cb1])
-  show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))))
-          \<le> e * content (cbox (u,w) (v,z))"
-    using content_cbox_pair_eq0_D [OF c0']
-    by (force simp add: c0')
-next
-  fix a::'a and c::'b and b::'a and d::'b
-  and M::real and i::'a and j::'b
-  and u::'a and v::'a and w::'b and z::'b
-  assume ij: "(i,j) \<in> Basis"
-     and n1: "\<forall>a' b' c' d'.
-                cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
-                cbox (a',c') (b',d') \<subseteq> {x. x \<bullet> (i,j) \<le> M} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
-                norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
-                \<le> e * content (cbox (a',c') (b',d'))"
-     and n2: "\<forall>a' b' c' d'.
-                cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and>
-                cbox (a',c') (b',d') \<subseteq> {x. M \<le> x \<bullet> (i,j)} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
-                norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y))))
-                \<le> e * content (cbox (a',c') (b',d'))"
-     and subs: "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)"  "cbox (u,w) (v,z) \<subseteq> s"
-  have fcont: "continuous_on (cbox (u, w) (v, z)) f"
-    using assms(1) continuous_on_subset  subs(2) by blast
-  then have fint: "f integrable_on cbox (u, w) (v, z)"
-    by (metis integrable_continuous)
-  consider "i \<in> Basis" "j=0" | "j \<in> Basis" "i=0"  using ij
-    by (auto simp: Euclidean_Space.Basis_prod_def)
-  then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
-             \<le> e * content (cbox (u,w) (v,z))" (is ?normle)
-  proof cases
-    case 1
-    then have e: "e * content (cbox (u, w) (v, z)) =
-                  e * (content (cbox u v \<inter> {x. x \<bullet> i \<le> M}) * content (cbox w z)) +
-                  e * (content (cbox u v \<inter> {x. M \<le> x \<bullet> i}) * content (cbox w z))"
-      by (simp add: content_split [where c=M] content_Pair algebra_simps)
-    have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
-                integral (cbox u v \<inter> {x. x \<bullet> i \<le> M}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) +
-                integral (cbox u v \<inter> {x. M \<le> x \<bullet> i}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))"
-      using 1 f subs integral_integrable_2dim continuous_on_subset
-      by (blast intro: integral_split)
-    show ?normle
-      apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
-      using 1 subs
-      apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\<lambda>u. M\<le>u"] setcomp_dot1 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp1)
-      apply (simp_all add: interval_split ij)
-      apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
-      apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
-      apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
-      done
-  next
-    case 2
-    then have e: "e * content (cbox (u, w) (v, z)) =
-                  e * (content (cbox u v) * content (cbox w z \<inter> {x. x \<bullet> j \<le> M})) +
-                  e * (content (cbox u v) * content (cbox w z \<inter> {x. M \<le> x \<bullet> j}))"
-      by (simp add: content_split [where c=M] content_Pair algebra_simps)
-    have "(\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
-                "(\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
-      using 2 subs
-      apply (simp_all add: interval_split)
-      apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]])
-      apply (auto simp: cbox_Pair_eq interval_split [symmetric])
-      done
-    with 2 have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) =
-                   integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) +
-                   integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y)))"
-      by (simp add: integral_add [symmetric] integral_split [symmetric]
-                    continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong)
-    show ?normle
-      apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e])
-      using 2 subs
-      apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\<lambda>u. M\<le>u"] setcomp_dot2 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp2)
-      apply (simp_all add: interval_split ij)
-      apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric])
-      apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format])
-      apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format])
-      done
-  qed
-qed
-
-lemma integral_Pair_const:
-    "integral (cbox (a,c) (b,d)) (\<lambda>x. k) =
-     integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. k))"
-  by (simp add: content_Pair)
-
-lemma norm_minus2: "norm (x1-x2, y1-y2) = norm (x2-x1, y2-y1)"
-  by (simp add: norm_minus_eqI)
-
-lemma integral_prod_continuous:
-  fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
-  assumes "continuous_on (cbox (a,c) (b,d)) f" (is "continuous_on ?CBOX f")
-    shows "integral (cbox (a,c) (b,d)) f = integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f(x,y)))"
-proof (cases "content ?CBOX = 0")
-  case True
-  then show ?thesis
-    by (auto simp: content_Pair)
-next
-  case False
-  then have cbp: "content ?CBOX > 0"
-    using content_lt_nz by blast
-  have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) = 0"
-  proof (rule dense_eq0_I, simp)
-    fix e::real  assume "0 < e"
-    with cbp have e': "0 < e / content ?CBOX"
-      by simp
-    have f_int_acbd: "f integrable_on cbox (a,c) (b,d)"
-      by (rule integrable_continuous [OF assms])
-    { fix p
-      assume p: "p division_of cbox (a,c) (b,d)"
-      note opd1 = comm_monoid_set.operative_division [OF comm_monoid_set_and integral_swap_operative [OF assms e'], THEN iffD1,
-                       THEN spec, THEN spec, THEN spec, THEN spec, of p "(a,c)" "(b,d)" a c b d]
-      have "(\<And>t u v w z.
-              \<lbrakk>t \<in> p; cbox (u,w) (v,z) \<subseteq> t;  cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)\<rbrakk> \<Longrightarrow>
-              norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
-              \<le> e * content (cbox (u,w) (v,z)) / content?CBOX)
-            \<Longrightarrow>
-            norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
-        using opd1 [OF p] False  by (simp add: comm_monoid_set_F_and)
-    } note op_acbd = this
-    { fix k::real and p and u::'a and v w and z::'b and t1 t2 l
-      assume k: "0 < k"
-         and nf: "\<And>x y u v.
-                  \<lbrakk>x \<in> cbox a b; y \<in> cbox c d; u \<in> cbox a b; v\<in>cbox c d; norm (u-x, v-y) < k\<rbrakk>
-                  \<Longrightarrow> norm (f(u,v) - f(x,y)) < e / (2 * (content ?CBOX))"
-         and p_acbd: "p tagged_division_of cbox (a,c) (b,d)"
-         and fine: "(\<lambda>x. ball x k) fine p"  "((t1,t2), l) \<in> p"
-         and uwvz_sub: "cbox (u,w) (v,z) \<subseteq> l" "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)"
-      have t: "t1 \<in> cbox a b" "t2 \<in> cbox c d"
-        by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+
-      have ls: "l \<subseteq> ball (t1,t2) k"
-        using fine by (simp add: fine_def Ball_def)
-      { fix x1 x2
-        assume xuvwz: "x1 \<in> cbox u v" "x2 \<in> cbox w z"
-        then have x: "x1 \<in> cbox a b" "x2 \<in> cbox c d"
-          using uwvz_sub by auto
-        have "norm (x1 - t1, x2 - t2) < k"
-          using xuvwz ls uwvz_sub unfolding ball_def
-          by (force simp add: cbox_Pair_eq dist_norm norm_minus2)
-        then have "norm (f (x1,x2) - f (t1,t2)) \<le> e / content ?CBOX / 2"
-          using nf [OF t x]  by simp
-      } note nf' = this
-      have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)"
-        using f_int_acbd uwvz_sub integrable_on_subcbox by blast
-      have f_int_uv: "\<And>x. x \<in> cbox u v \<Longrightarrow> (\<lambda>y. f (x,y)) integrable_on cbox w z"
-        using assms continuous_on_subset uwvz_sub
-        by (blast intro: continuous_on_imp_integrable_on_Pair1)
-      have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\<lambda>x. f (t1,t2)))
-         \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
-        apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const)
-        apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]])
-        using cbp e' nf'
-        apply (auto simp: integrable_diff f_int_uwvz integrable_const)
-        done
-      have int_integrable: "(\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) integrable_on cbox u v"
-        using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast
-      have normint_wz:
-         "\<And>x. x \<in> cbox u v \<Longrightarrow>
-               norm (integral (cbox w z) (\<lambda>y. f (x, y)) - integral (cbox w z) (\<lambda>y. f (t1, t2)))
-               \<le> e * content (cbox w z) / content (cbox (a, c) (b, d)) / 2"
-        apply (simp only: integral_diff [symmetric] f_int_uv integrable_const)
-        apply (rule order_trans [OF integrable_bound [of "e / content ?CBOX / 2"]])
-        using cbp e' nf'
-        apply (auto simp: integrable_diff f_int_uv integrable_const)
-        done
-      have "norm (integral (cbox u v)
-               (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)) - integral (cbox w z) (\<lambda>y. f (t1,t2))))
-            \<le> e * content (cbox w z) / content ?CBOX / 2 * content (cbox u v)"
-        apply (rule integrable_bound [OF _ _ normint_wz])
-        using cbp e'
-        apply (auto simp: divide_simps content_pos_le integrable_diff int_integrable integrable_const)
-        done
-      also have "... \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
-        by (simp add: content_Pair divide_simps)
-      finally have 2: "norm (integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))) -
-                      integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (t1,t2))))
-                \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX / 2"
-        by (simp only: integral_diff [symmetric] int_integrable integrable_const)
-      have norm_xx: "\<lbrakk>x' = y'; norm(x - x') \<le> e/2; norm(y - y') \<le> e/2\<rbrakk> \<Longrightarrow> norm(x - y) \<le> e" for x::'c and y x' y' e
-        using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] real_sum_of_halves
-        by (simp add: norm_minus_commute)
-      have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))))
-            \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX"
-        by (rule norm_xx [OF integral_Pair_const 1 2])
-    } note * = this
-    show "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e"
-      using compact_uniformly_continuous [OF assms compact_cbox]
-      apply (simp add: uniformly_continuous_on_def dist_norm)
-      apply (drule_tac x="e / 2 / content?CBOX" in spec)
-      using cbp \<open>0 < e\<close>
-      apply (auto simp: zero_less_mult_iff)
-      apply (rename_tac k)
-      apply (rule_tac e1=k in fine_division_exists [OF gauge_ball, where a = "(a,c)" and b = "(b,d)"])
-      apply assumption
-      apply (rule op_acbd)
-      apply (erule division_of_tagged_division)
-      using *
-      apply auto
-      done
-  qed
-  then show ?thesis
-    by simp
-qed
-
-lemma swap_continuous:
-  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
-    shows "continuous_on (cbox (c,a) (d,b)) (\<lambda>(x, y). f y x)"
-proof -
-  have "(\<lambda>(x, y). f y x) = (\<lambda>(x, y). f x y) \<circ> prod.swap"
-    by auto
-  then show ?thesis
-    apply (rule ssubst)
-    apply (rule continuous_on_compose)
-    apply (simp add: split_def)
-    apply (rule continuous_intros | simp add: assms)+
-    done
-qed
-
-lemma integral_swap_2dim:
-  fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
-  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
-    shows "integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y) = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
-proof -
-  have "((\<lambda>(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))"
-    apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\<lambda>(x,y). f y x" "integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)", simplified])
-    apply (auto simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)
-    done
- then show ?thesis
-   by force
-qed
-
-theorem integral_swap_continuous:
-  fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach"
-  assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)"
-    shows "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) =
-           integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
-proof -
-  have "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y)"
-    using integral_prod_continuous [OF assms] by auto
-  also have "... = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)"
-    by (rule integral_swap_2dim [OF assms])
-  also have "... = integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))"
-    using integral_prod_continuous [OF swap_continuous] assms
-    by auto
-  finally show ?thesis .
-qed
-
-
-subsection \<open>Definite integrals for exponential and power function\<close>
-
-lemma has_integral_exp_minus_to_infinity:
-  assumes a: "a > 0"
-  shows   "((\<lambda>x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}"
-proof -
-  define f where "f = (\<lambda>k x. if x \<in> {c..real k} then exp (-a*x) else 0)"
-
-  {
-    fix k :: nat assume k: "of_nat k \<ge> c"
-    from k a
-      have "((\<lambda>x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}"
-      by (intro fundamental_theorem_of_calculus)
-         (auto intro!: derivative_eq_intros
-               simp: has_field_derivative_iff_has_vector_derivative [symmetric])
-    hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def
-      by (subst has_integral_restrict) simp_all
-  } note has_integral_f = this
-
-  have [simp]: "f k = (\<lambda>_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def)
-  have integral_f: "integral {c..} (f k) =
-                      (if real k \<ge> c then exp (-a*c)/a - exp (-a*real k)/a else 0)"
-    for k using integral_unique[OF has_integral_f[of k]] by simp
-
-  have A: "(\<lambda>x. exp (-a*x)) integrable_on {c..} \<and>
-             (\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> integral {c..} (\<lambda>x. exp (-a*x))"
-  proof (intro monotone_convergence_increasing allI ballI)
-    fix k ::nat
-    have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P)
-      unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real)
-    hence int: "(f k) integrable_on {c..of_real k}"
-      by (rule integrable_eq[rotated]) (simp add: f_def)
-    show "f k integrable_on {c..}"
-      by (rule integrable_on_superset[OF _ _ int]) (auto simp: f_def)
-  next
-    fix x assume x: "x \<in> {c..}"
-    have "sequentially \<le> principal {nat \<lceil>x\<rceil>..}" unfolding at_top_def by (simp add: Inf_lower)
-    also have "{nat \<lceil>x\<rceil>..} \<subseteq> {k. x \<le> real k}" by auto
-    also have "inf (principal \<dots>) (principal {k. \<not>x \<le> real k}) =
-                 principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" by simp
-    also have "{k. x \<le> real k} \<inter> {k. \<not>x \<le> real k} = {}" by blast
-    finally have "inf sequentially (principal {k. \<not>x \<le> real k}) = bot"
-      by (simp add: inf.coboundedI1 bot_unique)
-    with x show "(\<lambda>k. f k x) \<longlonglongrightarrow> exp (-a*x)" unfolding f_def
-      by (intro filterlim_If) auto
-  next
-    have "\<bar>integral {c..} (f k)\<bar> \<le> exp (-a*c)/a" for k
-    proof (cases "c > of_nat k")
-      case False
-      hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)"
-        by (simp add: integral_f)
-      also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) =
-                   exp (- (a * c)) / a - exp (- (a * real k)) / a"
-        using False a by (intro abs_of_nonneg) (simp_all add: field_simps)
-      also have "\<dots> \<le> exp (- a * c) / a" using a by simp
-      finally show ?thesis .
-    qed (insert a, simp_all add: integral_f)
-    thus "bounded {integral {c..} (f k) |k. True}"
-      by (intro bounded_realI[of _ "exp (-a*c)/a"]) auto
-  qed (auto simp: f_def)
-
-  from eventually_gt_at_top[of "nat \<lceil>c\<rceil>"] have "eventually (\<lambda>k. of_nat k > c) sequentially"
-    by eventually_elim linarith
-  hence "eventually (\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially"
-    by eventually_elim (simp add: integral_f)
-  moreover have "(\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
-    by (intro tendsto_intros filterlim_compose[OF exp_at_bot]
-          filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+
-       (insert a, simp_all)
-  ultimately have "(\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
-    by (rule Lim_transform_eventually)
-  from LIMSEQ_unique[OF conjunct2[OF A] this]
-    have "integral {c..} (\<lambda>x. exp (-a*x)) = exp (-a*c)/a" by simp
-  with conjunct1[OF A] show ?thesis
-    by (simp add: has_integral_integral)
-qed
-
-lemma integrable_on_exp_minus_to_infinity: "a > 0 \<Longrightarrow> (\<lambda>x. exp (-a*x) :: real) integrable_on {c..}"
-  using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast
-
-lemma has_integral_powr_from_0:
-  assumes a: "a > (-1::real)" and c: "c \<ge> 0"
-  shows   "((\<lambda>x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}"
-proof (cases "c = 0")
-  case False
-  define f where "f = (\<lambda>k x. if x \<in> {inverse (of_nat (Suc k))..c} then x powr a else 0)"
-  define F where "F = (\<lambda>k. if inverse (of_nat (Suc k)) \<le> c then
-                             c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)"
-
-  {
-    fix k :: nat
-    have "(f k has_integral F k) {0..c}"
-    proof (cases "inverse (of_nat (Suc k)) \<le> c")
-      case True
-      {
-        fix x assume x: "x \<ge> inverse (1 + real k)"
-        have "0 < inverse (1 + real k)" by simp
-        also note x
-        finally have "x > 0" .
-      } note x = this
-      hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) -
-               inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}"
-        using True a by (intro fundamental_theorem_of_calculus)
-           (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const
-             continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric])
-      with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all
-    next
-      case False
-      thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto
-    qed
-  } note has_integral_f = this
-  have integral_f: "integral {0..c} (f k) = F k" for k
-    using has_integral_f[of k] by (rule integral_unique)
-
-  have A: "(\<lambda>x. x powr a) integrable_on {0..c} \<and>
-           (\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> integral {0..c} (\<lambda>x. x powr a)"
-  proof (intro monotone_convergence_increasing ballI allI)
-    fix k from has_integral_f[of k] show "f k integrable_on {0..c}"
-      by (auto simp: integrable_on_def)
-  next
-    fix k :: nat and x :: real
-    {
-      assume x: "inverse (real (Suc k)) \<le> x"
-      have "inverse (real (Suc (Suc k))) \<le> inverse (real (Suc k))" by (simp add: field_simps)
-      also note x
-      finally have "inverse (real (Suc (Suc k))) \<le> x" .
-    }
-    thus "f k x \<le> f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc)
-  next
-    fix x assume x: "x \<in> {0..c}"
-    show "(\<lambda>k. f k x) \<longlonglongrightarrow> x powr a"
-    proof (cases "x = 0")
-      case False
-      with x have "x > 0" by simp
-      from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
-        have "eventually (\<lambda>k. x powr a = f k x) sequentially"
-        by eventually_elim (insert x, simp add: f_def)
-      moreover have "(\<lambda>_. x powr a) \<longlonglongrightarrow> x powr a" by simp
-      ultimately show ?thesis by (rule Lim_transform_eventually)
-    qed (simp_all add: f_def)
-  next
-    {
-      fix k
-      from a have "F k \<le> c powr (a + 1) / (a + 1)"
-        by (auto simp add: F_def divide_simps)
-      also from a have "F k \<ge> 0"
-        by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2)
-      hence "F k = abs (F k)" by simp
-      finally have "abs (F k) \<le>  c powr (a + 1) / (a + 1)" .
-    }
-    thus "bounded {integral {0..c} (f k) |k. True}"
-      by (intro bounded_realI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f)
-  qed
-
-  from False c have "c > 0" by simp
-  from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this]
-    have "eventually (\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) =
-            integral {0..c} (f k)) sequentially"
-    by eventually_elim (simp add: integral_f F_def)
-  moreover have "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
-                   \<longlonglongrightarrow> c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)"
-    using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto
-  hence "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1))
-          \<longlonglongrightarrow> c powr (a + 1) / (a + 1)" by simp
-  ultimately have "(\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> c powr (a+1) / (a+1)"
-    by (rule Lim_transform_eventually)
-  with A have "integral {0..c} (\<lambda>x. x powr a) = c powr (a+1) / (a+1)"
-    by (blast intro: LIMSEQ_unique)
-  with A show ?thesis by (simp add: has_integral_integral)
-qed (simp_all add: has_integral_refl)
-
-lemma integrable_on_powr_from_0:
-  assumes a: "a > (-1::real)" and c: "c \<ge> 0"
-  shows   "(\<lambda>x. x powr a) integrable_on {0..c}"
-  using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast
-
-end
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -1,16 +1,17 @@
 theory Multivariate_Analysis
 imports
-  Fashoda
+  Fashoda_Theorem
   Extended_Real_Limits
   Determinants
   Homeomorphism
   Ordered_Euclidean_Space
   Bounded_Continuous_Function
-  Weierstrass
+  Weierstrass_Theorems
   Polytope
+  Poly_Roots
   Conformal_Mappings
   Generalised_Binomial_Theorem
-  Gamma
+  Gamma_Function
 begin
 
 end
--- a/src/HOL/Multivariate_Analysis/Path_Connected.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Path_Connected.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -5,7 +5,7 @@
 section \<open>Continuous paths and path-connected sets\<close>
 
 theory Path_Connected
-imports Extension
+imports Continuous_Extension
 begin
 
 subsection \<open>Paths and Arcs\<close>
@@ -558,9 +558,9 @@
 
 subsection\<open>Some reversed and "if and only if" versions of joining theorems\<close>
 
-lemma path_join_path_ends: 
+lemma path_join_path_ends:
   fixes g1 :: "real \<Rightarrow> 'a::metric_space"
-  assumes "path(g1 +++ g2)" "path g2" 
+  assumes "path(g1 +++ g2)" "path g2"
     shows "pathfinish g1 = pathstart g2"
 proof (rule ccontr)
   define e where "e = dist (g1 1) (g2 0)"
@@ -568,14 +568,14 @@
   then have "0 < dist (pathfinish g1) (pathstart g2)"
     by auto
   then have "e > 0"
-    by (metis e_def pathfinish_def pathstart_def) 
-  then obtain d1 where "d1 > 0" 
+    by (metis e_def pathfinish_def pathstart_def)
+  then obtain d1 where "d1 > 0"
        and d1: "\<And>x'. \<lbrakk>x'\<in>{0..1}; norm x' < d1\<rbrakk> \<Longrightarrow> dist (g2 x') (g2 0) < e/2"
     using assms(2) unfolding path_def continuous_on_iff
     apply (drule_tac x=0 in bspec, simp)
     by (metis half_gt_zero_iff norm_conv_dist)
-  obtain d2 where "d2 > 0" 
-       and d2: "\<And>x'. \<lbrakk>x'\<in>{0..1}; dist x' (1/2) < d2\<rbrakk> 
+  obtain d2 where "d2 > 0"
+       and d2: "\<And>x'. \<lbrakk>x'\<in>{0..1}; dist x' (1/2) < d2\<rbrakk>
                       \<Longrightarrow> dist ((g1 +++ g2) x') (g1 1) < e/2"
     using assms(1) \<open>e > 0\<close> unfolding path_def continuous_on_iff
     apply (drule_tac x="1/2" in bspec, simp)
@@ -597,25 +597,25 @@
     using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def)
   then have "dist (g1 1) (g2 0) < e/2 + e/2"
     using dist_triangle_half_r e_def by blast
-  then show False 
+  then show False
     by (simp add: e_def [symmetric])
 qed
 
-lemma path_join_eq [simp]:  
+lemma path_join_eq [simp]:
   fixes g1 :: "real \<Rightarrow> 'a::metric_space"
   assumes "path g1" "path g2"
     shows "path(g1 +++ g2) \<longleftrightarrow> pathfinish g1 = pathstart g2"
   using assms by (metis path_join_path_ends path_join_imp)
 
-lemma simple_path_joinE: 
+lemma simple_path_joinE:
   assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2"
-  obtains "arc g1" "arc g2" 
+  obtains "arc g1" "arc g2"
           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 proof -
-  have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk> 
+  have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
                \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
     using assms by (simp add: simple_path_def)
-  have "path g1" 
+  have "path g1"
     using assms path_join simple_path_imp_path by blast
   moreover have "inj_on g1 {0..1}"
   proof (clarsimp simp: inj_on_def)
@@ -627,7 +627,7 @@
   ultimately have "arc g1"
     using assms  by (simp add: arc_def)
   have [simp]: "g2 0 = g1 1"
-    using assms by (metis pathfinish_def pathstart_def) 
+    using assms by (metis pathfinish_def pathstart_def)
   have "path g2"
     using assms path_join simple_path_imp_path by blast
   moreover have "inj_on g2 {0..1}"
@@ -640,7 +640,7 @@
   qed
   ultimately have "arc g2"
     using assms  by (simp add: arc_def)
-  have "g2 y = g1 0 \<or> g2 y = g1 1" 
+  have "g2 y = g1 0 \<or> g2 y = g1 1"
        if "g1 x = g2 y" "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" for x y
       using * [of "x / 2" "(y + 1) / 2"] that
       by (auto simp: joinpaths_def split_ifs divide_simps)
@@ -650,20 +650,20 @@
 qed
 
 lemma simple_path_join_loop_eq:
-  assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" 
+  assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2"
     shows "simple_path(g1 +++ g2) \<longleftrightarrow>
              arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
 by (metis assms simple_path_joinE simple_path_join_loop)
 
 lemma arc_join_eq:
-  assumes "pathfinish g1 = pathstart g2" 
+  assumes "pathfinish g1 = pathstart g2"
     shows "arc(g1 +++ g2) \<longleftrightarrow>
            arc g1 \<and> arc g2 \<and> path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
            (is "?lhs = ?rhs")
-proof 
+proof
   assume ?lhs
   then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path)
-  then have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk> 
+  then have *: "\<And>x y. \<lbrakk>0 \<le> x; x \<le> 1; 0 \<le> y; y \<le> 1; (g1 +++ g2) x = (g1 +++ g2) y\<rbrakk>
                \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
     using assms by (simp add: simple_path_def)
   have False if "g1 0 = g2 u" "0 \<le> u" "u \<le> 1" for u
@@ -681,7 +681,7 @@
     by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join)
 qed
 
-lemma arc_join_eq_alt: 
+lemma arc_join_eq_alt:
         "pathfinish g1 = pathstart g2
         \<Longrightarrow> (arc(g1 +++ g2) \<longleftrightarrow>
              arc g1 \<and> arc g2 \<and>
@@ -696,27 +696,27 @@
      \<Longrightarrow> path(p +++ (q +++ r)) \<longleftrightarrow> path((p +++ q) +++ r)"
 by simp
 
-lemma simple_path_assoc: 
-  assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" 
+lemma simple_path_assoc:
+  assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r"
     shows "simple_path (p +++ (q +++ r)) \<longleftrightarrow> simple_path ((p +++ q) +++ r)"
 proof (cases "pathstart p = pathfinish r")
   case True show ?thesis
   proof
     assume "simple_path (p +++ q +++ r)"
     with assms True show "simple_path ((p +++ q) +++ r)"
-      by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join 
+      by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join
                     dest: arc_distinct_ends [of r])
   next
     assume 0: "simple_path ((p +++ q) +++ r)"
     with assms True have q: "pathfinish r \<notin> path_image q"
-      using arc_distinct_ends  
+      using arc_distinct_ends
       by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join)
     have "pathstart r \<notin> path_image p"
       using assms
-      by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff 
+      by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff
               pathfinish_in_path_image pathfinish_join simple_path_joinE)
     with assms 0 q True show "simple_path (p +++ q +++ r)"
-      by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join 
+      by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join
                dest!: subsetD [OF _ IntI])
   qed
 next
@@ -730,11 +730,11 @@
     with a have "x = pathstart q"
       by blast
   }
-  with False assms show ?thesis 
+  with False assms show ?thesis
     by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join)
 qed
 
-lemma arc_assoc: 
+lemma arc_assoc:
      "\<lbrakk>pathfinish p = pathstart q; pathfinish q = pathstart r\<rbrakk>
       \<Longrightarrow> arc(p +++ (q +++ r)) \<longleftrightarrow> arc((p +++ q) +++ r)"
 by (simp add: arc_simple_path simple_path_assoc)
@@ -1262,7 +1262,7 @@
   shows "\<exists>e > 0. ball z e \<inter> path_image g = {}"
 proof -
   obtain a where "a \<in> path_image g" "\<forall>y \<in> path_image g. dist z a \<le> dist z y"
-    apply (rule distance_attains_inf[OF _ path_image_nonempty, of g z]) 
+    apply (rule distance_attains_inf[OF _ path_image_nonempty, of g z])
     using compact_path_image[THEN compact_imp_closed, OF assms(1)] by auto
   then show ?thesis
     apply (rule_tac x="dist z a" in exI)
@@ -2080,8 +2080,8 @@
 corollary connected_open_delete_finite:
   fixes S T::"'a::euclidean_space set"
   assumes S: "open S" "connected S" and 2: "2 \<le> DIM('a)" and "finite T"
-  shows "connected(S - T)" 
-  using \<open>finite T\<close> S 
+  shows "connected(S - T)"
+  using \<open>finite T\<close> S
 proof (induct T)
   case empty
   show ?case using \<open>connected S\<close> by simp
@@ -3068,12 +3068,12 @@
       by (force simp: imageI image_mono interiorI interior_subset frontier_def y)
     have **: "(~(b \<inter> (- S) = {}) \<and> ~(b - (- S) = {}) \<Longrightarrow> (b \<inter> f \<noteq> {}))
               \<Longrightarrow> (b \<inter> S \<noteq> {}) \<Longrightarrow> b \<inter> f = {} \<Longrightarrow>
-              b \<subseteq> S" for b f and S :: "'b set" 
+              b \<subseteq> S" for b f and S :: "'b set"
       by blast
     show ?thesis
       apply (rule **)   (*such a horrible mess*)
       apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball])
-      using \<open>a \<in> S\<close> \<open>0 < r\<close> 
+      using \<open>a \<in> S\<close> \<open>0 < r\<close>
       apply (auto simp: disjoint_iff_not_equal  dist_norm)
       by (metis dw_le norm_minus_commute not_less order_trans rle wy)
 qed
@@ -4003,7 +4003,7 @@
 lemma homotopic_points_eq_path_component:
    "homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow>
         path_component S a b"
-by (auto simp: path_component_imp_homotopic_points 
+by (auto simp: path_component_imp_homotopic_points
          dest: homotopic_loops_imp_path_component_value [where t=1])
 
 lemma path_connected_eq_homotopic_points:
@@ -4058,11 +4058,11 @@
  apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any)
 apply (clarsimp simp add: simply_connected_eq_contractible_loop_any)
 apply (drule_tac x=p in spec)
-using homotopic_loops_trans path_connected_eq_homotopic_points 
+using homotopic_loops_trans path_connected_eq_homotopic_points
   apply blast
 done
 
-lemma simply_connected_eq_contractible_loop_all: 
+lemma simply_connected_eq_contractible_loop_all:
   fixes S :: "_::real_normed_vector set"
   shows "simply_connected S \<longleftrightarrow>
          S = {} \<or>
@@ -4075,21 +4075,21 @@
   case False
   then obtain a where "a \<in> S" by blast
   show ?thesis
-  proof  
+  proof
     assume "simply_connected S"
     then show ?rhs
-      using \<open>a \<in> S\<close> \<open>simply_connected S\<close> simply_connected_eq_contractible_loop_any 
+      using \<open>a \<in> S\<close> \<open>simply_connected S\<close> simply_connected_eq_contractible_loop_any
       by blast
-  next     
+  next
     assume ?rhs
     then show "simply_connected S"
       apply (simp add: simply_connected_eq_contractible_loop_any False)
-      by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans 
+      by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans
              path_component_imp_homotopic_points path_component_refl)
   qed
 qed
 
-lemma simply_connected_eq_contractible_path: 
+lemma simply_connected_eq_contractible_path:
   fixes S :: "_::real_normed_vector set"
   shows "simply_connected S \<longleftrightarrow>
            path_connected S \<and>
@@ -4098,7 +4098,7 @@
 apply (rule iffI)
  apply (simp add: simply_connected_imp_path_connected)
  apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null)
-by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image 
+by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image
          simply_connected_eq_contractible_loop_some subset_iff)
 
 lemma simply_connected_eq_homotopic_paths:
@@ -4112,23 +4112,23 @@
          (is "?lhs = ?rhs")
 proof
   assume ?lhs
-  then have pc: "path_connected S" 
+  then have pc: "path_connected S"
         and *:  "\<And>p. \<lbrakk>path p; path_image p \<subseteq> S;
-                       pathfinish p = pathstart p\<rbrakk> 
+                       pathfinish p = pathstart p\<rbrakk>
                       \<Longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p))"
     by (auto simp: simply_connected_eq_contractible_path)
-  have "homotopic_paths S p q" 
+  have "homotopic_paths S p q"
         if "path p" "path_image p \<subseteq> S" "path q"
            "path_image q \<subseteq> S" "pathstart q = pathstart p"
            "pathfinish q = pathfinish p" for p q
   proof -
-    have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" 
+    have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))"
       by (simp add: homotopic_paths_rid homotopic_paths_sym that)
     also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p))
                                  (p +++ reversepath q +++ q)"
       using that
       by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath)
-    also have "homotopic_paths S (p +++ reversepath q +++ q) 
+    also have "homotopic_paths S (p +++ reversepath q +++ q)
                                  ((p +++ reversepath q) +++ q)"
       by (simp add: that homotopic_paths_assoc)
     also have "homotopic_paths S ((p +++ reversepath q) +++ q)
@@ -4142,7 +4142,7 @@
   then show ?rhs
     by (blast intro: pc *)
 next
-  assume ?rhs 
+  assume ?rhs
   then show ?lhs
     by (force simp: simply_connected_eq_contractible_path)
 qed
@@ -4408,7 +4408,7 @@
 lemma is_interval_contractible_1:
   fixes S :: "real set"
   shows  "is_interval S \<longleftrightarrow> contractible S"
-using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 
+using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1
       is_interval_simply_connected_1 by auto
 
 lemma contractible_Times:
@@ -4416,14 +4416,14 @@
   assumes S: "contractible S" and T: "contractible T"
   shows "contractible (S \<times> T)"
 proof -
-  obtain a h where conth: "continuous_on ({0..1} \<times> S) h" 
+  obtain a h where conth: "continuous_on ({0..1} \<times> S) h"
              and hsub: "h ` ({0..1} \<times> S) \<subseteq> S"
-             and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x" 
+             and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x"
              and [simp]: "\<And>x. x \<in> S \<Longrightarrow>  h (1::real, x) = a"
     using S by (auto simp add: contractible_def homotopic_with)
-  obtain b k where contk: "continuous_on ({0..1} \<times> T) k" 
+  obtain b k where contk: "continuous_on ({0..1} \<times> T) k"
              and ksub: "k ` ({0..1} \<times> T) \<subseteq> T"
-             and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x" 
+             and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x"
              and [simp]: "\<And>x. x \<in> T \<Longrightarrow>  k (1::real, x) = b"
     using T by (auto simp add: contractible_def homotopic_with)
   show ?thesis
@@ -4432,16 +4432,16 @@
     apply (rule exI [where x=b])
     apply (rule exI [where x = "\<lambda>z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"])
     apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk])
-    using hsub ksub 
+    using hsub ksub
     apply (auto simp: )
     done
 qed
 
-lemma homotopy_dominated_contractibility: 
+lemma homotopy_dominated_contractibility:
   fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
   assumes S: "contractible S"
-      and f: "continuous_on S f" "image f S \<subseteq> T" 
-      and g: "continuous_on T g" "image g T \<subseteq> S" 
+      and f: "continuous_on S f" "image f S \<subseteq> T"
+      and g: "continuous_on T g" "image g T \<subseteq> S"
       and hom: "homotopic_with (\<lambda>x. True) T T (f o g) id"
     shows "contractible T"
 proof -
--- a/src/HOL/Multivariate_Analysis/PolyRoots.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,301 +0,0 @@
-(*  Author: John Harrison and Valentina Bruno
-    Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson
-*)
-
-section \<open>polynomial functions: extremal behaviour and root counts\<close>
-
-theory PolyRoots
-imports Complex_Main
-begin
-
-subsection\<open>Geometric progressions\<close>
-
-lemma setsum_gp_basic:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  shows "(1 - x) * (\<Sum>i\<le>n. x^i) = 1 - x^Suc n"
-  by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost)
-
-lemma setsum_gp0:
-  fixes x :: "'a::{comm_ring,division_ring}"
-  shows   "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))"
-  using setsum_gp_basic[of x n]
-  by (simp add: mult.commute divide_simps)
-
-lemma setsum_power_add:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  shows "(\<Sum>i\<in>I. x^(m+i)) = x^m * (\<Sum>i\<in>I. x^i)"
-  by (simp add: setsum_right_distrib power_add)
-
-lemma setsum_power_shift:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  assumes "m \<le> n"
-  shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
-proof -
-  have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
-    by (simp add: setsum_right_distrib power_add [symmetric])
-  also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)"
-    using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
-  finally show ?thesis .
-qed
-
-lemma setsum_gp_multiplied:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  assumes "m \<le> n"
-  shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n"
-proof -
-  have  "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)"
-    by (metis mult.assoc mult.commute assms setsum_power_shift)
-  also have "... =x^m * (1 - x^Suc(n-m))"
-    by (metis mult.assoc setsum_gp_basic)
-  also have "... = x^m - x^Suc n"
-    using assms
-    by (simp add: algebra_simps) (metis le_add_diff_inverse power_add)
-  finally show ?thesis .
-qed
-
-lemma setsum_gp:
-  fixes x :: "'a::{comm_ring,division_ring}"
-  shows   "(\<Sum>i=m..n. x^i) =
-               (if n < m then 0
-                else if x = 1 then of_nat((n + 1) - m)
-                else (x^m - x^Suc n) / (1 - x))"
-using setsum_gp_multiplied [of m n x]
-apply auto
-by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq)
-
-lemma setsum_gp_offset:
-  fixes x :: "'a::{comm_ring,division_ring}"
-  shows   "(\<Sum>i=m..m+n. x^i) =
-       (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
-  using setsum_gp [of x m "m+n"]
-  by (auto simp: power_add algebra_simps)
-
-lemma setsum_gp_strict:
-  fixes x :: "'a::{comm_ring,division_ring}"
-  shows "(\<Sum>i<n. x^i) = (if x = 1 then of_nat n else (1 - x^n) / (1 - x))"
-  by (induct n) (auto simp: algebra_simps divide_simps)
-
-subsection\<open>Basics about polynomial functions: extremal behaviour and root counts.\<close>
-
-lemma sub_polyfun:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
-           (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)"
-proof -
-  have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
-        (\<Sum>i\<le>n. a i * (x^i - y^i))"
-    by (simp add: algebra_simps setsum_subtractf [symmetric])
-  also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
-    by (simp add: power_diff_sumr2 ac_simps)
-  also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
-    by (simp add: setsum_right_distrib ac_simps)
-  also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
-    by (simp add: nested_setsum_swap')
-  finally show ?thesis .
-qed
-
-lemma sub_polyfun_alt:
-  fixes x :: "'a::{comm_ring,monoid_mult}"
-  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
-           (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)"
-proof -
-  { fix j
-    have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =
-          (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"
-      by (rule setsum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto }
-  then show ?thesis
-    by (simp add: sub_polyfun)
-qed
-
-lemma polyfun_linear_factor:
-  fixes a :: "'a::{comm_ring,monoid_mult}"
-  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) =
-                  (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)"
-proof -
-  { fix z
-    have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) =
-          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
-      by (simp add: sub_polyfun setsum_left_distrib)
-    then have "(\<Sum>i\<le>n. c i * z^i) =
-          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
-          + (\<Sum>i\<le>n. c i * a^i)"
-      by (simp add: algebra_simps) }
-  then show ?thesis
-    by (intro exI allI)
-qed
-
-lemma polyfun_linear_factor_root:
-  fixes a :: "'a::{comm_ring,monoid_mult}"
-  assumes "(\<Sum>i\<le>n. c i * a^i) = 0"
-  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)"
-  using polyfun_linear_factor [of c n a] assms
-  by simp
-
-lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b"
-  by (metis norm_triangle_mono order.trans order_refl)
-
-lemma polyfun_extremal_lemma:
-  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
-  assumes "e > 0"
-    shows "\<exists>M. \<forall>z. M \<le> norm z \<longrightarrow> norm(\<Sum>i\<le>n. c i * z^i) \<le> e * norm(z) ^ Suc n"
-proof (induction n)
-  case 0
-  show ?case
-    by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms)
-next
-  case (Suc n)
-  then obtain M where M: "\<forall>z. M \<le> norm z \<longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" ..
-  show ?case
-  proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify)
-    fix z::'a
-    assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z"
-    then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z"
-      by auto
-    then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z"  "(norm z * norm z ^ n) > 0"
-      apply (metis assms less_divide_eq mult.commute not_le)
-      using norm1 apply (metis mult_pos_pos zero_less_power)
-      done
-    have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) =
-          (e + norm (c (Suc n))) * (norm z * norm z ^ n)"
-      by (simp add: norm_mult norm_power algebra_simps)
-    also have "... \<le> (e * norm z) * (norm z * norm z ^ n)"
-      using norm2 by (metis real_mult_le_cancel_iff1)
-    also have "... = e * (norm z * (norm z * norm z ^ n))"
-      by (simp add: algebra_simps)
-    finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n))
-                  \<le> e * (norm z * (norm z * norm z ^ n))" .
-    then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1
-      by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle)
-    qed
-qed
-
-lemma norm_lemma_xy: assumes "\<bar>b\<bar> + 1 \<le> norm(y) - a" "norm(x) \<le> a" shows "b \<le> norm(x + y)"
-proof -
-  have "b \<le> norm y - norm x"
-    using assms by linarith
-  then show ?thesis
-    by (metis (no_types) add.commute norm_diff_ineq order_trans)
-qed
-
-lemma polyfun_extremal:
-  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
-  assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0"
-    shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity"
-using assms
-proof (induction n)
-  case 0 then show ?case
-    by simp
-next
-  case (Suc n)
-  show ?case
-  proof (cases "c (Suc n) = 0")
-    case True
-    with Suc show ?thesis
-      by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq)
-  next
-    case False
-    with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n]
-    obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow>
-               norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n"
-      by auto
-    show ?thesis
-    unfolding eventually_at_infinity
-    proof (rule exI [where x="max M (max 1 ((\<bar>B\<bar> + 1) / (norm (c (Suc n)) / 2)))"], clarsimp)
-      fix z::'a
-      assume les: "M \<le> norm z"  "1 \<le> norm z"  "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z"
-      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"
-        by (metis False pos_divide_le_eq zero_less_norm_iff)
-      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))"
-        by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
-      then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les
-        apply auto
-        apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])
-        apply (simp_all add: norm_mult norm_power)
-        done
-    qed
-  qed
-qed
-
-lemma polyfun_rootbound:
- fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
- assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
-   shows "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<and> card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
-using assms
-proof (induction n arbitrary: c)
- case (Suc n) show ?case
- proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
-   case False
-   then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0"
-     by auto
-   from polyfun_linear_factor_root [OF this]
-   obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)"
-     by auto
-   then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)"
-     by (metis lessThan_Suc_atMost)
-   then have ins_ab: "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = insert a {z. (\<Sum>i\<le>n. b i * z^i) = 0}"
-     by auto
-   have c0: "c 0 = - (a * b 0)" using  b [of 0]
-     by simp
-   then have extr_prem: "~ (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"
-     by (metis Suc.prems le0 minus_zero mult_zero_right)
-   have "\<exists>k\<le>n. b k \<noteq> 0"
-     apply (rule ccontr)
-     using polyfun_extremal [OF extr_prem, of 1]
-     apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc)
-     apply (drule_tac x="of_real ba" in spec, simp)
-     done
-   then show ?thesis using Suc.IH [of b] ins_ab
-     by (auto simp: card_insert_if)
-   qed simp
-qed simp
-
-corollary
-  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
-  assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
-    shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
-      and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
-using polyfun_rootbound [OF assms] by auto
-
-lemma polyfun_finite_roots:
-  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
-    shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
-proof (cases " \<exists>k\<le>n. c k \<noteq> 0")
-  case True then show ?thesis
-    by (blast intro: polyfun_rootbound_finite)
-next
-  case False then show ?thesis
-    by (auto simp: infinite_UNIV_char_0)
-qed
-
-lemma polyfun_eq_0:
-  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
-    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"
-proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")
-  case True
-  then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
-    by (simp add: infinite_UNIV_char_0)
-  with True show ?thesis
-    by (metis (poly_guards_query) polyfun_rootbound_finite)
-next
-  case False
-  then show ?thesis
-    by auto
-qed
-
-lemma polyfun_eq_const:
-  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
-    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
-proof -
-  {fix z
-    have "(\<Sum>i\<le>n. c i * z^i) = (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) + k"
-      by (induct n) auto
-  } then
-  have "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> (\<forall>z. (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) = 0)"
-    by auto
-  also have "... \<longleftrightarrow>  c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
-    by (auto simp: polyfun_eq_0)
-  finally show ?thesis .
-qed
-
-end
-
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Poly_Roots.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,301 @@
+(*  Author: John Harrison and Valentina Bruno
+    Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson
+*)
+
+section \<open>polynomial functions: extremal behaviour and root counts\<close>
+
+theory Poly_Roots
+imports Complex_Main
+begin
+
+subsection\<open>Geometric progressions\<close>
+
+lemma setsum_gp_basic:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows "(1 - x) * (\<Sum>i\<le>n. x^i) = 1 - x^Suc n"
+  by (simp only: one_diff_power_eq [of "Suc n" x] lessThan_Suc_atMost)
+
+lemma setsum_gp0:
+  fixes x :: "'a::{comm_ring,division_ring}"
+  shows   "(\<Sum>i\<le>n. x^i) = (if x = 1 then of_nat(n + 1) else (1 - x^Suc n) / (1 - x))"
+  using setsum_gp_basic[of x n]
+  by (simp add: mult.commute divide_simps)
+
+lemma setsum_power_add:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows "(\<Sum>i\<in>I. x^(m+i)) = x^m * (\<Sum>i\<in>I. x^i)"
+  by (simp add: setsum_right_distrib power_add)
+
+lemma setsum_power_shift:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  assumes "m \<le> n"
+  shows "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i\<le>n-m. x^i)"
+proof -
+  have "(\<Sum>i=m..n. x^i) = x^m * (\<Sum>i=m..n. x^(i-m))"
+    by (simp add: setsum_right_distrib power_add [symmetric])
+  also have "(\<Sum>i=m..n. x^(i-m)) = (\<Sum>i\<le>n-m. x^i)"
+    using \<open>m \<le> n\<close> by (intro setsum.reindex_bij_witness[where j="\<lambda>i. i - m" and i="\<lambda>i. i + m"]) auto
+  finally show ?thesis .
+qed
+
+lemma setsum_gp_multiplied:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  assumes "m \<le> n"
+  shows "(1 - x) * (\<Sum>i=m..n. x^i) = x^m - x^Suc n"
+proof -
+  have  "(1 - x) * (\<Sum>i=m..n. x^i) = x^m * (1 - x) * (\<Sum>i\<le>n-m. x^i)"
+    by (metis mult.assoc mult.commute assms setsum_power_shift)
+  also have "... =x^m * (1 - x^Suc(n-m))"
+    by (metis mult.assoc setsum_gp_basic)
+  also have "... = x^m - x^Suc n"
+    using assms
+    by (simp add: algebra_simps) (metis le_add_diff_inverse power_add)
+  finally show ?thesis .
+qed
+
+lemma setsum_gp:
+  fixes x :: "'a::{comm_ring,division_ring}"
+  shows   "(\<Sum>i=m..n. x^i) =
+               (if n < m then 0
+                else if x = 1 then of_nat((n + 1) - m)
+                else (x^m - x^Suc n) / (1 - x))"
+using setsum_gp_multiplied [of m n x]
+apply auto
+by (metis eq_iff_diff_eq_0 mult.commute nonzero_divide_eq_eq)
+
+lemma setsum_gp_offset:
+  fixes x :: "'a::{comm_ring,division_ring}"
+  shows   "(\<Sum>i=m..m+n. x^i) =
+       (if x = 1 then of_nat n + 1 else x^m * (1 - x^Suc n) / (1 - x))"
+  using setsum_gp [of x m "m+n"]
+  by (auto simp: power_add algebra_simps)
+
+lemma setsum_gp_strict:
+  fixes x :: "'a::{comm_ring,division_ring}"
+  shows "(\<Sum>i<n. x^i) = (if x = 1 then of_nat n else (1 - x^n) / (1 - x))"
+  by (induct n) (auto simp: algebra_simps divide_simps)
+
+subsection\<open>Basics about polynomial functions: extremal behaviour and root counts.\<close>
+
+lemma sub_polyfun:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
+           (x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)"
+proof -
+  have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
+        (\<Sum>i\<le>n. a i * (x^i - y^i))"
+    by (simp add: algebra_simps setsum_subtractf [symmetric])
+  also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"
+    by (simp add: power_diff_sumr2 ac_simps)
+  also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"
+    by (simp add: setsum_right_distrib ac_simps)
+  also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"
+    by (simp add: nested_setsum_swap')
+  finally show ?thesis .
+qed
+
+lemma sub_polyfun_alt:
+  fixes x :: "'a::{comm_ring,monoid_mult}"
+  shows   "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =
+           (x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)"
+proof -
+  { fix j
+    have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =
+          (\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"
+      by (rule setsum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto }
+  then show ?thesis
+    by (simp add: sub_polyfun)
+qed
+
+lemma polyfun_linear_factor:
+  fixes a :: "'a::{comm_ring,monoid_mult}"
+  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) =
+                  (z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)"
+proof -
+  { fix z
+    have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) =
+          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"
+      by (simp add: sub_polyfun setsum_left_distrib)
+    then have "(\<Sum>i\<le>n. c i * z^i) =
+          (z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)
+          + (\<Sum>i\<le>n. c i * a^i)"
+      by (simp add: algebra_simps) }
+  then show ?thesis
+    by (intro exI allI)
+qed
+
+lemma polyfun_linear_factor_root:
+  fixes a :: "'a::{comm_ring,monoid_mult}"
+  assumes "(\<Sum>i\<le>n. c i * a^i) = 0"
+  shows  "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)"
+  using polyfun_linear_factor [of c n a] assms
+  by simp
+
+lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b"
+  by (metis norm_triangle_mono order.trans order_refl)
+
+lemma polyfun_extremal_lemma:
+  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
+  assumes "e > 0"
+    shows "\<exists>M. \<forall>z. M \<le> norm z \<longrightarrow> norm(\<Sum>i\<le>n. c i * z^i) \<le> e * norm(z) ^ Suc n"
+proof (induction n)
+  case 0
+  show ?case
+    by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms)
+next
+  case (Suc n)
+  then obtain M where M: "\<forall>z. M \<le> norm z \<longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" ..
+  show ?case
+  proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify)
+    fix z::'a
+    assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z"
+    then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z"
+      by auto
+    then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z"  "(norm z * norm z ^ n) > 0"
+      apply (metis assms less_divide_eq mult.commute not_le)
+      using norm1 apply (metis mult_pos_pos zero_less_power)
+      done
+    have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) =
+          (e + norm (c (Suc n))) * (norm z * norm z ^ n)"
+      by (simp add: norm_mult norm_power algebra_simps)
+    also have "... \<le> (e * norm z) * (norm z * norm z ^ n)"
+      using norm2 by (metis real_mult_le_cancel_iff1)
+    also have "... = e * (norm z * (norm z * norm z ^ n))"
+      by (simp add: algebra_simps)
+    finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n))
+                  \<le> e * (norm z * (norm z * norm z ^ n))" .
+    then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1
+      by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle)
+    qed
+qed
+
+lemma norm_lemma_xy: assumes "\<bar>b\<bar> + 1 \<le> norm(y) - a" "norm(x) \<le> a" shows "b \<le> norm(x + y)"
+proof -
+  have "b \<le> norm y - norm x"
+    using assms by linarith
+  then show ?thesis
+    by (metis (no_types) add.commute norm_diff_ineq order_trans)
+qed
+
+lemma polyfun_extremal:
+  fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"
+  assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0"
+    shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity"
+using assms
+proof (induction n)
+  case 0 then show ?case
+    by simp
+next
+  case (Suc n)
+  show ?case
+  proof (cases "c (Suc n) = 0")
+    case True
+    with Suc show ?thesis
+      by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq)
+  next
+    case False
+    with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n]
+    obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow>
+               norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n"
+      by auto
+    show ?thesis
+    unfolding eventually_at_infinity
+    proof (rule exI [where x="max M (max 1 ((\<bar>B\<bar> + 1) / (norm (c (Suc n)) / 2)))"], clarsimp)
+      fix z::'a
+      assume les: "M \<le> norm z"  "1 \<le> norm z"  "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z"
+      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"
+        by (metis False pos_divide_le_eq zero_less_norm_iff)
+      then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))"
+        by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)
+      then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les
+        apply auto
+        apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])
+        apply (simp_all add: norm_mult norm_power)
+        done
+    qed
+  qed
+qed
+
+lemma polyfun_rootbound:
+ fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+ assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
+   shows "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<and> card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
+using assms
+proof (induction n arbitrary: c)
+ case (Suc n) show ?case
+ proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
+   case False
+   then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0"
+     by auto
+   from polyfun_linear_factor_root [OF this]
+   obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)"
+     by auto
+   then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)"
+     by (metis lessThan_Suc_atMost)
+   then have ins_ab: "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = insert a {z. (\<Sum>i\<le>n. b i * z^i) = 0}"
+     by auto
+   have c0: "c 0 = - (a * b 0)" using  b [of 0]
+     by simp
+   then have extr_prem: "~ (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"
+     by (metis Suc.prems le0 minus_zero mult_zero_right)
+   have "\<exists>k\<le>n. b k \<noteq> 0"
+     apply (rule ccontr)
+     using polyfun_extremal [OF extr_prem, of 1]
+     apply (auto simp: eventually_at_infinity b simp del: setsum_atMost_Suc)
+     apply (drule_tac x="of_real ba" in spec, simp)
+     done
+   then show ?thesis using Suc.IH [of b] ins_ab
+     by (auto simp: card_insert_if)
+   qed simp
+qed simp
+
+corollary
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+  assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"
+    shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
+      and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
+using polyfun_rootbound [OF assms] by auto
+
+lemma polyfun_finite_roots:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
+proof (cases " \<exists>k\<le>n. c k \<noteq> 0")
+  case True then show ?thesis
+    by (blast intro: polyfun_rootbound_finite)
+next
+  case False then show ?thesis
+    by (auto simp: infinite_UNIV_char_0)
+qed
+
+lemma polyfun_eq_0:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"
+proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")
+  case True
+  then have "~ finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
+    by (simp add: infinite_UNIV_char_0)
+  with True show ?thesis
+    by (metis (poly_guards_query) polyfun_rootbound_finite)
+next
+  case False
+  then show ?thesis
+    by auto
+qed
+
+lemma polyfun_eq_const:
+  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
+    shows  "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
+proof -
+  {fix z
+    have "(\<Sum>i\<le>n. c i * z^i) = (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) + k"
+      by (induct n) auto
+  } then
+  have "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> (\<forall>z. (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) = 0)"
+    by auto
+  also have "... \<longleftrightarrow>  c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"
+    by (auto simp: polyfun_eq_0)
+  finally show ?thesis .
+qed
+
+end
+
--- a/src/HOL/Multivariate_Analysis/Summation.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,920 +0,0 @@
-(*  Title:    HOL/Multivariate_Analysis/Summation.thy
-    Author:   Manuel Eberl, TU München
-*)
-  
-section \<open>Radius of Convergence and Summation Tests\<close>
-
-theory Summation
-imports
-  Complex_Main
-  "~~/src/HOL/Library/Extended_Real" 
-  "~~/src/HOL/Library/Liminf_Limsup"
-begin
-
-text \<open>
-  The definition of the radius of convergence of a power series, 
-  various summability tests, lemmas to compute the radius of convergence etc.
-\<close>
-
-subsection \<open>Rounded dual logarithm\<close>
-
-(* This is required for the Cauchy condensation criterion *)
-
-definition "natlog2 n = (if n = 0 then 0 else nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
-
-lemma natlog2_0 [simp]: "natlog2 0 = 0" by (simp add: natlog2_def)
-lemma natlog2_1 [simp]: "natlog2 1 = 0" by (simp add: natlog2_def)
-lemma natlog2_eq_0_iff: "natlog2 n = 0 \<longleftrightarrow> n < 2" by (simp add: natlog2_def)
-
-lemma natlog2_power_of_two [simp]: "natlog2 (2 ^ n) = n"
-  by (simp add: natlog2_def log_nat_power)
-
-lemma natlog2_mono: "m \<le> n \<Longrightarrow> natlog2 m \<le> natlog2 n"
-  unfolding natlog2_def by (simp_all add: nat_mono floor_mono)
-
-lemma pow_natlog2_le: "n > 0 \<Longrightarrow> 2 ^ natlog2 n \<le> n"
-proof -
-  assume n: "n > 0"
-  from n have "of_nat (2 ^ natlog2 n) = 2 powr real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
-    by (subst powr_realpow) (simp_all add: natlog2_def)
-  also have "\<dots> = 2 powr of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>" using n by simp
-  also have "\<dots> \<le> 2 powr log 2 (real_of_nat n)" by (intro powr_mono) (linarith, simp_all)
-  also have "\<dots> = of_nat n" using n by simp
-  finally show ?thesis by simp
-qed
-
-lemma pow_natlog2_gt: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n > n"
-  and pow_natlog2_ge: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n \<ge> n"
-proof -
-  assume n: "n > 0"
-  from n have "of_nat n = 2 powr log 2 (real_of_nat n)" by simp
-  also have "\<dots> < 2 powr (1 + of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>)" 
-    by (intro powr_less_mono) (linarith, simp_all)
-  also from n have "\<dots> = 2 powr (1 + real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>))" by simp
-  also from n have "\<dots> = of_nat (2 * 2 ^ natlog2 n)"
-    by (simp_all add: natlog2_def powr_real_of_int powr_add)
-  finally show "2 * 2 ^ natlog2 n > n" by (rule of_nat_less_imp_less)
-  thus "2 * 2 ^ natlog2 n \<ge> n" by simp
-qed
-
-lemma natlog2_eqI:
-  assumes "n > 0" "2^k \<le> n" "n < 2 * 2^k"
-  shows   "natlog2 n = k"
-proof -
-  from assms have "of_nat (2 ^ k) \<le> real_of_nat n"  by (subst of_nat_le_iff) simp_all
-  hence "real_of_int (int k) \<le> log (of_nat 2) (real_of_nat n)"
-    by (subst le_log_iff) (simp_all add: powr_realpow assms del: of_nat_le_iff)
-  moreover from assms have "real_of_nat n < of_nat (2 ^ Suc k)" by (subst of_nat_less_iff) simp_all
-  hence "log 2 (real_of_nat n) < of_nat k + 1"
-    by (subst log_less_iff) (simp_all add: assms powr_realpow powr_add)
-  ultimately have "\<lfloor>log 2 (real_of_nat n)\<rfloor> = of_nat k" by (intro floor_unique) simp_all
-  with assms show ?thesis by (simp add: natlog2_def)
-qed
-
-lemma natlog2_rec: 
-  assumes "n \<ge> 2"
-  shows   "natlog2 n = 1 + natlog2 (n div 2)"
-proof (rule natlog2_eqI)
-  from assms have "2 ^ (1 + natlog2 (n div 2)) \<le> 2 * (n div 2)" 
-    by (simp add: pow_natlog2_le)
-  also have "\<dots> \<le> n" by simp
-  finally show "2 ^ (1 + natlog2 (n div 2)) \<le> n" .
-next
-  from assms have "n < 2 * (n div 2 + 1)" by simp 
-  also from assms have "(n div 2) < 2 ^ (1 + natlog2 (n div 2))" 
-    by (simp add: pow_natlog2_gt)
-  hence "2 * (n div 2 + 1) \<le> 2 * (2 ^ (1 + natlog2 (n div 2)))" 
-    by (intro mult_left_mono) simp_all
-  finally show "n < 2 * 2 ^ (1 + natlog2 (n div 2))" .
-qed (insert assms, simp_all)
-
-fun natlog2_aux where
-  "natlog2_aux n acc = (if (n::nat) < 2 then acc else natlog2_aux (n div 2) (acc + 1))"
-
-lemma natlog2_aux_correct:
-  "natlog2_aux n acc = acc + natlog2 n"
-  by (induction n acc rule: natlog2_aux.induct) (auto simp: natlog2_rec natlog2_eq_0_iff)
-  
-lemma natlog2_code [code]: "natlog2 n = natlog2_aux n 0"
-  by (subst natlog2_aux_correct) simp
-
-
-subsection \<open>Convergence tests for infinite sums\<close>
-
-subsubsection \<open>Root test\<close>
-
-lemma limsup_root_powser:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  shows "limsup (\<lambda>n. ereal (root n (norm (f n * z ^ n)))) = 
-             limsup (\<lambda>n. ereal (root n (norm (f n)))) * ereal (norm z)"
-proof -
-  have A: "(\<lambda>n. ereal (root n (norm (f n * z ^ n)))) = 
-              (\<lambda>n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h")
-  proof
-    fix n show "?g n = ?h n"
-    by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power)
-  qed
-  show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all
-qed
-
-lemma limsup_root_limit:
-  assumes "(\<lambda>n. ereal (root n (norm (f n)))) \<longlonglongrightarrow> l" (is "?g \<longlonglongrightarrow> _")
-  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = l"
-proof -
-  from assms have "convergent ?g" "lim ?g = l"
-    unfolding convergent_def by (blast intro: limI)+
-  with convergent_limsup_cl show ?thesis by force
-qed
-
-lemma limsup_root_limit':
-  assumes "(\<lambda>n. root n (norm (f n))) \<longlonglongrightarrow> l"
-  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = ereal l"
-  by (intro limsup_root_limit tendsto_ereal assms)
-
-lemma root_test_convergence':
-  fixes f :: "nat \<Rightarrow> 'a :: banach"
-  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
-  assumes l: "l < 1"
-  shows   "summable f"
-proof -
-  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
-  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
-  finally have "l \<ge> 0" by simp
-  with l obtain l' where l': "l = ereal l'" by (cases l) simp_all
-
-  define c where "c = (1 - l') / 2"
-  from l and \<open>l \<ge> 0\<close> have c: "l + c > l" "l' + c \<ge> 0" "l' + c < 1" unfolding c_def 
-    by (simp_all add: field_simps l')
-  have "\<forall>C>l. eventually (\<lambda>n. ereal (root n (norm (f n))) < C) sequentially"
-    by (subst Limsup_le_iff[symmetric]) (simp add: l_def)
-  with c have "eventually (\<lambda>n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp
-  with eventually_gt_at_top[of "0::nat"]
-    have "eventually (\<lambda>n. norm (f n) \<le> (l' + c) ^ n) sequentially"
-  proof eventually_elim
-    fix n :: nat assume n: "n > 0" 
-    assume "ereal (root n (norm (f n))) < l + ereal c"
-    hence "root n (norm (f n)) \<le> l' + c" by (simp add: l')
-    with c n have "root n (norm (f n)) ^ n \<le> (l' + c) ^ n"
-      by (intro power_mono) (simp_all add: real_root_ge_zero)
-    also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp
-    finally show "norm (f n) \<le> (l' + c) ^ n" by simp
-  qed
-  thus ?thesis
-    by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c)
-qed
-
-lemma root_test_divergence:
-  fixes f :: "nat \<Rightarrow> 'a :: banach"
-  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
-  assumes l: "l > 1"
-  shows   "\<not>summable f"
-proof
-  assume "summable f"
-  hence bounded: "Bseq f" by (simp add: summable_imp_Bseq)
-
-  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
-  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
-  finally have l_nonneg: "l \<ge> 0" by simp
-
-  define c where "c = (if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2)"
-  from l l_nonneg consider "l = \<infinity>" | "\<exists>l'. l = ereal l'" by (cases l) simp_all
-  hence c: "c > 1 \<and> ereal c < l" by cases (insert l, auto simp: c_def field_simps)
-
-  have unbounded: "\<not>bdd_above {n. root n (norm (f n)) > c}"
-  proof
-    assume "bdd_above {n. root n (norm (f n)) > c}"
-    then obtain N where "\<forall>n. root n (norm (f n)) > c \<longrightarrow> n \<le> N" unfolding bdd_above_def by blast
-    hence "\<exists>N. \<forall>n\<ge>N. root n (norm (f n)) \<le> c"
-      by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric])
-    hence "eventually (\<lambda>n. root n (norm (f n)) \<le> c) sequentially"
-      by (auto simp: eventually_at_top_linorder)
-    hence "l \<le> c" unfolding l_def by (intro Limsup_bounded) simp_all
-    with c show False by auto
-  qed
-  
-  from bounded obtain K where K: "K > 0" "\<And>n. norm (f n) \<le> K" using BseqE by blast
-  define n where "n = nat \<lceil>log c K\<rceil>"
-  from unbounded have "\<exists>m>n. c < root m (norm (f m))" unfolding bdd_above_def
-    by (auto simp: not_le)
-  then guess m by (elim exE conjE) note m = this
-  from c K have "K = c powr log c K" by (simp add: powr_def log_def)
-  also from c have "c powr log c K \<le> c powr real n" unfolding n_def
-    by (intro powr_mono, linarith, simp)
-  finally have "K \<le> c ^ n" using c by (simp add: powr_realpow)
-  also from c m have "c ^ n < c ^ m" by simp
-  also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all
-  also from m have "... = norm (f m)" by simp
-  finally show False using K(2)[of m]  by simp
-qed
-
-
-subsubsection \<open>Cauchy's condensation test\<close>
-
-context
-fixes f :: "nat \<Rightarrow> real"
-begin
-
-private lemma condensation_inequality:
-  assumes mono: "\<And>m n. 0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> f n \<le> f m"
-  shows   "(\<Sum>k=1..<n. f k) \<ge> (\<Sum>k=1..<n. f (2 * 2 ^ natlog2 k))" (is "?thesis1")
-          "(\<Sum>k=1..<n. f k) \<le> (\<Sum>k=1..<n. f (2 ^ natlog2 k))" (is "?thesis2")
-  by (intro setsum_mono mono pow_natlog2_ge pow_natlog2_le, simp, simp)+
-
-private lemma condensation_condense1: "(\<Sum>k=1..<2^n. f (2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ k))"
-proof (induction n)
-  case (Suc n)
-  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto  
-  also have "(\<Sum>k\<in>\<dots>. f (2 ^ natlog2 k)) = 
-                 (\<Sum>k<n. 2^k * f (2^k)) + (\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k))" 
-    by (subst setsum.union_disjoint) (insert Suc, auto)
-  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
-  hence "(\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
-    by (intro setsum.cong) simp_all
-  also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
-  finally show ?case by simp
-qed simp
-
-private lemma condensation_condense2: "(\<Sum>k=1..<2^n. f (2 * 2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ Suc k))"
-proof (induction n)
-  case (Suc n)
-  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto  
-  also have "(\<Sum>k\<in>\<dots>. f (2 * 2 ^ natlog2 k)) = 
-                 (\<Sum>k<n. 2^k * f (2^Suc k)) + (\<Sum>k = 2^n..<2^Suc n. f (2 * 2^natlog2 k))" 
-    by (subst setsum.union_disjoint) (insert Suc, auto)
-  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
-  hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
-    by (intro setsum.cong) simp_all
-  also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
-  finally show ?case by simp
-qed simp
-
-lemma condensation_test:
-  assumes mono: "\<And>m. 0 < m \<Longrightarrow> f (Suc m) \<le> f m"
-  assumes nonneg: "\<And>n. f n \<ge> 0"
-  shows "summable f \<longleftrightarrow> summable (\<lambda>n. 2^n * f (2^n))"
-proof -
-  define f' where "f' n = (if n = 0 then 0 else f n)" for n
-  from mono have mono': "decseq (\<lambda>n. f (Suc n))" by (intro decseq_SucI) simp
-  hence mono': "f n \<le> f m" if "m \<le> n" "m > 0" for m n 
-    using that decseqD[OF mono', of "m - 1" "n - 1"] by simp
-  
-  have "(\<lambda>n. f (Suc n)) = (\<lambda>n. f' (Suc n))" by (intro ext) (simp add: f'_def)
-  hence "summable f \<longleftrightarrow> summable f'"
-    by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:)
-  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k<n. f' k)" unfolding summable_iff_convergent ..
-  also have "monoseq (\<lambda>n. \<Sum>k<n. f' k)" unfolding f'_def
-    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
-  hence "convergent (\<lambda>n. \<Sum>k<n. f' k) \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. f' k)"
-    by (rule monoseq_imp_convergent_iff_Bseq)
-  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f' k)" unfolding One_nat_def
-    by (subst setsum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
-  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f k)" unfolding f'_def by simp
-  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
-    by (rule nonneg_incseq_Bseq_subseq_iff[symmetric])
-       (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def)
-  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^k))"
-  proof (intro iffI)
-    assume A: "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
-    have "eventually (\<lambda>n. norm (\<Sum>k<n. 2^k * f (2^Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)) sequentially"
-    proof (intro always_eventually allI)
-      fix n :: nat
-      have "norm (\<Sum>k<n. 2^k * f (2^Suc k)) = (\<Sum>k<n. 2^k * f (2^Suc k))" unfolding real_norm_def
-        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
-      also have "\<dots> \<le> (\<Sum>k=1..<2^n. f k)"
-        by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono')
-      also have "\<dots> = norm \<dots>" unfolding real_norm_def
-        by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg)
-      finally show "norm (\<Sum>k<n. 2 ^ k * f (2 ^ Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)" .
-    qed
-    from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
-    from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
-      by (simp add: setsum_right_distrib setsum_left_distrib mult_ac)
-    hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
-      by (intro Bseq_add, subst setsum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
-    hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
-      by (subst setsum_head_upt_Suc) (simp_all add: add_ac)
-    thus "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))" 
-      by (subst (asm) Bseq_Suc_iff) (simp add: atLeast0LessThan)
-  next
-    assume A: "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
-    have "eventually (\<lambda>n. norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))) sequentially"
-    proof (intro always_eventually allI)
-      fix n :: nat
-      have "norm (\<Sum>k=1..<2^n. f k) = (\<Sum>k=1..<2^n. f k)" unfolding real_norm_def
-        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg)
-      also have "\<dots> \<le> (\<Sum>k<n. 2^k * f (2^k))"
-        by (subst condensation_condense1 [symmetric]) (intro condensation_inequality mono')
-      also have "\<dots> = norm \<dots>" unfolding real_norm_def
-        by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
-      finally show "norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))" .
-    qed
-    from this and A show "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)" by (rule Bseq_eventually_mono)
-  qed
-  also have "monoseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
-    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
-  hence "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k))) \<longleftrightarrow> convergent (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
-    by (rule monoseq_imp_convergent_iff_Bseq [symmetric])
-  also have "\<dots> \<longleftrightarrow> summable (\<lambda>k. 2^k * f (2^k))" by (simp only: summable_iff_convergent)
-  finally show ?thesis .
-qed
-
-end
-
-
-subsubsection \<open>Summability of powers\<close>
-
-lemma abs_summable_complex_powr_iff: 
-    "summable (\<lambda>n. norm (exp (of_real (ln (of_nat n)) * s))) \<longleftrightarrow> Re s < -1"
-proof (cases "Re s \<le> 0")
-  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
-  case False
-  with eventually_gt_at_top[of "0::nat"]
-    have "eventually (\<lambda>n. norm (1 :: real) \<le> norm (exp (?l n * s))) sequentially" 
-    by (auto intro!: ge_one_powr_ge_zero elim!: eventually_mono)
-  from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff)
-next
-  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
-  case True
-  hence "summable (\<lambda>n. norm (exp (?l n * s))) \<longleftrightarrow> summable (\<lambda>n. 2^n * norm (exp (?l (2^n) * s)))"
-    by (intro condensation_test) (auto intro!: mult_right_mono_neg)
-  also have "(\<lambda>n. 2^n * norm (exp (?l (2^n) * s))) = (\<lambda>n. (2 powr (Re s + 1)) ^ n)"
-  proof
-    fix n :: nat
-    have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)"
-      using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps) 
-    also have "\<dots> = exp (real n * (ln 2 * (Re s + 1)))"
-      by (simp add: algebra_simps exp_add)
-    also have "\<dots> = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp
-    also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def)
-    finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" .
-  qed
-  also have "summable \<dots> \<longleftrightarrow> 2 powr (Re s + 1) < 2 powr 0"
-    by (subst summable_geometric_iff) simp
-  also have "\<dots> \<longleftrightarrow> Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith)
-  finally show ?thesis .
-qed
-
-lemma summable_complex_powr_iff: 
-  assumes "Re s < -1"
-  shows   "summable (\<lambda>n. exp (of_real (ln (of_nat n)) * s))"
-  by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact
-
-lemma summable_real_powr_iff: "summable (\<lambda>n. of_nat n powr s :: real) \<longleftrightarrow> s < -1"
-proof -
-  from eventually_gt_at_top[of "0::nat"]
-    have "summable (\<lambda>n. of_nat n powr s) \<longleftrightarrow> summable (\<lambda>n. exp (ln (of_nat n) * s))"
-    by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def)
-  also have "\<dots> \<longleftrightarrow> s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp
-  finally show ?thesis .
-qed
-
-lemma inverse_power_summable:
-  assumes s: "s \<ge> 2"
-  shows "summable (\<lambda>n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))"
-proof (rule summable_norm_cancel, subst summable_cong)
-  from eventually_gt_at_top[of "0::nat"]
-    show "eventually (\<lambda>n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top"
-    by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow)
-qed (insert s summable_real_powr_iff[of "-s"], simp_all)
-
-lemma not_summable_harmonic: "\<not>summable (\<lambda>n. inverse (of_nat n) :: 'a :: real_normed_field)"
-proof
-  assume "summable (\<lambda>n. inverse (of_nat n) :: 'a)"
-  hence "convergent (\<lambda>n. norm (of_real (\<Sum>k<n. inverse (of_nat k)) :: 'a))" 
-    by (simp add: summable_iff_convergent convergent_norm)
-  hence "convergent (\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real)" by (simp only: norm_of_real)
-  also have "(\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real) = (\<lambda>n. \<Sum>k<n. inverse (of_nat k))"
-    by (intro ext abs_of_nonneg setsum_nonneg) auto
-  also have "convergent \<dots> \<longleftrightarrow> summable (\<lambda>k. inverse (of_nat k) :: real)"
-    by (simp add: summable_iff_convergent)
-  finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus)
-qed
-
-
-subsubsection \<open>Kummer's test\<close>
-
-lemma kummers_test_convergence:
-  fixes f p :: "nat \<Rightarrow> real"
-  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
-  assumes nonneg_p: "eventually (\<lambda>n. p n \<ge> 0) sequentially"
-  defines "l \<equiv> liminf (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
-  assumes l: "l > 0"
-  shows   "summable f"
-  unfolding summable_iff_convergent'
-proof -
-  define r where "r = (if l = \<infinity> then 1 else real_of_ereal l / 2)"
-  from l have "r > 0 \<and> of_real r < l" by (cases l) (simp_all add: r_def)
-  hence r: "r > 0" "of_real r < l" by simp_all
-  hence "eventually (\<lambda>n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially"
-    unfolding l_def by (force dest: less_LiminfD)
-  moreover from pos_f have "eventually (\<lambda>n. f (Suc n) > 0) sequentially" 
-    by (subst eventually_sequentially_Suc)
-  ultimately have "eventually (\<lambda>n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially"
-    by eventually_elim (simp add: field_simps)
-  from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]]
-    obtain m where m: "\<And>n. n \<ge> m \<Longrightarrow> f n > 0" "\<And>n. n \<ge> m \<Longrightarrow> p n \<ge> 0"
-        "\<And>n. n \<ge> m \<Longrightarrow> p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)"
-    unfolding eventually_at_top_linorder by blast
-
-  let ?c = "(norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r"
-  have "Bseq (\<lambda>n. (\<Sum>k\<le>n + Suc m. f k))"
-  proof (rule BseqI')
-    fix k :: nat
-    define n where "n = k + Suc m"
-    have n: "n > m" by (simp add: n_def)
-
-    from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
-      by (simp add: setsum_right_distrib[symmetric] abs_mult)
-    also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
-    hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
-    also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
-      by (subst setsum.union_disjoint) auto
-    also have "norm \<dots> \<le> norm (\<Sum>k\<le>m. r * f k) + norm (\<Sum>k=Suc m..n. r * f k)"
-      by (rule norm_triangle_ineq)
-    also from r less_imp_le[OF m(1)] have "(\<Sum>k=Suc m..n. r * f k) \<ge> 0" 
-      by (intro setsum_nonneg) auto
-    hence "norm (\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=Suc m..n. r * f k)" by simp
-    also have "(\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=m..<n. r * f (Suc k))"
-     by (subst setsum_shift_bounds_Suc_ivl [symmetric])
-          (simp only: atLeastLessThanSuc_atLeastAtMost)
-    also from m have "\<dots> \<le> (\<Sum>k=m..<n. p k * f k - p (Suc k) * f (Suc k))"
-      by (intro setsum_mono[OF less_imp_le]) simp_all
-    also have "\<dots> = -(\<Sum>k=m..<n. p (Suc k) * f (Suc k) - p k * f k)"
-      by (simp add: setsum_negf [symmetric] algebra_simps)
-    also from n have "\<dots> = p m * f m - p n * f n"
-      by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all
-    also from less_imp_le[OF m(1)] m(2) n have "\<dots> \<le> p m * f m" by simp
-    finally show "norm (\<Sum>k\<le>n. f k) \<le> (norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r" using r
-      by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac)
-  qed
-  moreover have "(\<Sum>k\<le>n. f k) \<le> (\<Sum>k\<le>n'. f k)" if "Suc m \<le> n" "n \<le> n'" for n n'
-    using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto
-  ultimately show "convergent (\<lambda>n. \<Sum>k\<le>n. f k)" by (rule Bseq_monoseq_convergent'_inc)
-qed
-
-
-lemma kummers_test_divergence:
-  fixes f p :: "nat \<Rightarrow> real"
-  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
-  assumes pos_p: "eventually (\<lambda>n. p n > 0) sequentially"
-  assumes divergent_p: "\<not>summable (\<lambda>n. inverse (p n))"
-  defines "l \<equiv> limsup (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
-  assumes l: "l < 0"
-  shows   "\<not>summable f"
-proof
-  assume "summable f"
-  from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]]
-    obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> p n > 0" "\<And>n. n \<ge> N \<Longrightarrow> f n > 0"
-                      "\<And>n. n \<ge> N \<Longrightarrow> p n * f n / f (Suc n) - p (Suc n) < 0"
-    by (auto simp: eventually_at_top_linorder)
-  hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \<ge> N" for n using that N[of n] N[of "Suc n"] 
-    by (simp add: field_simps)
-  have "p n * f n \<ge> p N * f N" if "n \<ge> N" for n using that and A
-      by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans)
-  from eventually_ge_at_top[of N] N this
-    have "eventually (\<lambda>n. norm (p N * f N * inverse (p n)) \<le> f n) sequentially"
-    by (auto elim!: eventually_mono simp: field_simps abs_of_pos)
-  from this and \<open>summable f\<close> have "summable (\<lambda>n. p N * f N * inverse (p n))"
-    by (rule summable_comparison_test_ev)
-  from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl] 
-    have "summable (\<lambda>n. inverse (p n))" by (simp add: divide_simps)
-  with divergent_p show False by contradiction
-qed
-
-
-subsubsection \<open>Ratio test\<close>
-
-lemma ratio_test_convergence:
-  fixes f :: "nat \<Rightarrow> real"
-  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
-  defines "l \<equiv> liminf (\<lambda>n. ereal (f n / f (Suc n)))"
-  assumes l: "l > 1"
-  shows   "summable f"
-proof (rule kummers_test_convergence[OF pos_f])
-  note l
-  also have "l = liminf (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1" 
-    by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
-  finally show "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) > 0"
-    by (cases "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
-qed simp
-
-lemma ratio_test_divergence:
-  fixes f :: "nat \<Rightarrow> real"
-  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
-  defines "l \<equiv> limsup (\<lambda>n. ereal (f n / f (Suc n)))"
-  assumes l: "l < 1"
-  shows   "\<not>summable f"
-proof (rule kummers_test_divergence[OF pos_f])
-  have "limsup (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1 = l" 
-    by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
-  also note l
-  finally show "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) < 0"
-    by (cases "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
-qed (simp_all add: summable_const_iff)
-
-
-subsubsection \<open>Raabe's test\<close>
-
-lemma raabes_test_convergence:
-fixes f :: "nat \<Rightarrow> real"
-  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
-  defines "l \<equiv> liminf (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
-  assumes l: "l > 1"
-  shows   "summable f"
-proof (rule kummers_test_convergence)
-  let ?l' = "liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
-  have "1 < l" by fact
-  also have "l = liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
-    by (simp add: l_def algebra_simps)
-  also have "\<dots> = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all
-  finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps)
-qed (simp_all add: pos)
-
-lemma raabes_test_divergence:
-fixes f :: "nat \<Rightarrow> real"
-  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
-  defines "l \<equiv> limsup (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
-  assumes l: "l < 1"
-  shows   "\<not>summable f"
-proof (rule kummers_test_divergence)
-  let ?l' = "limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
-  note l
-  also have "l = limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
-    by (simp add: l_def algebra_simps)
-  also have "\<dots> = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all
-  finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps)
-qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all)
-
-
-
-subsection \<open>Radius of convergence\<close>
-
-text \<open>
-  The radius of convergence of a power series. This value always exists, ranges from
-  @{term "0::ereal"} to @{term "\<infinity>::ereal"}, and the power series is guaranteed to converge for 
-  all inputs with a norm that is smaller than that radius and to diverge for all inputs with a
-  norm that is greater. 
-\<close>
-definition conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
-  "conv_radius f = inverse (limsup (\<lambda>n. ereal (root n (norm (f n)))))"
-
-lemma conv_radius_nonneg: "conv_radius f \<ge> 0"
-proof -
-  have "0 = limsup (\<lambda>n. 0)" by (subst Limsup_const) simp_all
-  also have "\<dots> \<le> limsup (\<lambda>n. ereal (root n (norm (f n))))"
-    by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
-  finally show ?thesis
-    unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff)
-qed
-
-lemma conv_radius_zero [simp]: "conv_radius (\<lambda>_. 0) = \<infinity>"
-  by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const)
-
-lemma conv_radius_cong:
-  assumes "eventually (\<lambda>x. f x = g x) sequentially"
-  shows   "conv_radius f = conv_radius g"
-proof -
-  have "eventually (\<lambda>n. ereal (root n (norm (f n))) = ereal (root n (norm (g n)))) sequentially"
-    using assms by eventually_elim simp
-  from Limsup_eq[OF this] show ?thesis unfolding conv_radius_def by simp
-qed
-
-lemma conv_radius_altdef:
-  "conv_radius f = liminf (\<lambda>n. inverse (ereal (root n (norm (f n)))))"
-  by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def)
-
-lemma abs_summable_in_conv_radius:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "ereal (norm z) < conv_radius f"
-  shows   "summable (\<lambda>n. norm (f n * z ^ n))"
-proof (rule root_test_convergence')
-  define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
-  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
-  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
-  finally have l_nonneg: "l \<ge> 0" .
-
-  have "limsup (\<lambda>n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def
-    by (rule limsup_root_powser)
-  also from l_nonneg consider "l = 0" | "l = \<infinity>" | "\<exists>l'. l = ereal l' \<and> l' > 0"
-    by (cases "l") (auto simp: less_le)
-  hence "l * ereal (norm z) < 1"
-  proof cases
-    assume "l = \<infinity>"
-    hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp
-    with assms show ?thesis by simp
-  next
-    assume "\<exists>l'. l = ereal l' \<and> l' > 0"
-    then guess l' by (elim exE conjE) note l' = this
-    hence "l \<noteq> \<infinity>" by auto
-    have "l * ereal (norm z) < l * conv_radius f"
-      by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms)
-    also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def)
-    also from l' have "l * inverse l = 1" by simp
-    finally show ?thesis .
-  qed simp_all
-  finally show "limsup (\<lambda>n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp
-qed
-
-lemma summable_in_conv_radius:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "ereal (norm z) < conv_radius f"
-  shows   "summable (\<lambda>n. f n * z ^ n)"
-  by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+
-
-lemma not_summable_outside_conv_radius:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "ereal (norm z) > conv_radius f"
-  shows   "\<not>summable (\<lambda>n. f n * z ^ n)"
-proof (rule root_test_divergence)
-  define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
-  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
-  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
-  finally have l_nonneg: "l \<ge> 0" .
-  from assms have l_nz: "l \<noteq> 0" unfolding conv_radius_def l_def by auto
-
-  have "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)"
-    unfolding l_def by (rule limsup_root_powser)
-  also have "... > 1"
-  proof (cases l)
-    assume "l = \<infinity>"
-    with assms conv_radius_nonneg[of f] show ?thesis
-      by (auto simp: zero_ereal_def[symmetric])
-  next
-    fix l' assume l': "l = ereal l'"
-    from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps)
-    also from l_nz have "inverse l = conv_radius f" 
-      unfolding l_def conv_radius_def by auto
-    also from l' l_nz l_nonneg assms have "l * \<dots> < l * ereal (norm z)"
-      by (intro ereal_mult_strict_left_mono) (auto simp: l')
-    finally show ?thesis .
-  qed (insert l_nonneg, simp_all)
-  finally show "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) > 1" .
-qed
-
-
-lemma conv_radius_geI:
-  assumes "summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
-  shows   "conv_radius f \<ge> norm z"
-  using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric])
-
-lemma conv_radius_leI:
-  assumes "\<not>summable (\<lambda>n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))"
-  shows   "conv_radius f \<le> norm z"
-  using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
-
-lemma conv_radius_leI':
-  assumes "\<not>summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
-  shows   "conv_radius f \<le> norm z"
-  using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
-
-lemma conv_radius_geI_ex:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
-  shows   "conv_radius f \<ge> R"
-proof (rule linorder_cases[of "conv_radius f" R])
-  assume R: "conv_radius f < R"
-  with conv_radius_nonneg[of f] obtain conv_radius' 
-    where [simp]: "conv_radius f = ereal conv_radius'"
-    by (cases "conv_radius f") simp_all
-  define r where "r = (if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)"
-  from R conv_radius_nonneg[of f] have "0 < r \<and> ereal r < R \<and> ereal r > conv_radius f" 
-    unfolding r_def by (cases R) (auto simp: r_def field_simps)
-  with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\<lambda>n. f n * z^n)" by auto
-  with not_summable_outside_conv_radius[of f z] show ?thesis by simp
-qed simp_all
-
-lemma conv_radius_geI_ex':
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * of_real r^n)"
-  shows   "conv_radius f \<ge> R"
-proof (rule conv_radius_geI_ex)
-  fix r assume "0 < r" "ereal r < R"
-  with assms[of r] show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)"
-    by (intro exI[of _ "of_real r :: 'a"]) auto
-qed
-
-lemma conv_radius_leI_ex:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "R \<ge> 0"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
-  shows   "conv_radius f \<le> R"
-proof (rule linorder_cases[of "conv_radius f" R])
-  assume R: "conv_radius f > R"
-  from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all
-  define r where
-    "r = (if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)"
-  from R conv_radius_nonneg[of f] have "r > R \<and> r < conv_radius f" unfolding r_def
-    by (cases "conv_radius f") (auto simp: r_def field_simps R')
-  with assms(1) assms(2)[of r] R' 
-    obtain z where "norm z < conv_radius f" "\<not>summable (\<lambda>n. norm (f n * z^n))" by auto
-  with abs_summable_in_conv_radius[of z f] show ?thesis by auto
-qed simp_all
-
-lemma conv_radius_leI_ex':
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "R \<ge> 0"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. f n * of_real r^n)"
-  shows   "conv_radius f \<le> R"
-proof (rule conv_radius_leI_ex)
-  fix r assume "0 < r" "ereal r > R"
-  with assms(2)[of r] show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))"
-    by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel)
-qed fact+
-
-lemma conv_radius_eqI:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "R \<ge> 0"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
-  shows   "conv_radius f = R"
-  by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms)
-
-lemma conv_radius_eqI':
-  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
-  assumes "R \<ge> 0"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * (of_real r)^n)"
-  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. norm (f n * (of_real r)^n))"
-  shows   "conv_radius f = R"
-proof (intro conv_radius_eqI[OF assms(1)])
-  fix r assume "0 < r" "ereal r < R" with assms(2)[OF this] 
-    show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)" by force
-next
-  fix r assume "0 < r" "ereal r > R" with assms(3)[OF this] 
-    show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))" by force  
-qed
-
-lemma conv_radius_zeroI:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "\<And>z. z \<noteq> 0 \<Longrightarrow> \<not>summable (\<lambda>n. f n * z^n)"
-  shows   "conv_radius f = 0"
-proof (rule ccontr)
-  assume "conv_radius f \<noteq> 0"
-  with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp
-  define r where "r = (if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2)"
-  from pos have r: "ereal r > 0 \<and> ereal r < conv_radius f" 
-    by (cases "conv_radius f") (simp_all add: r_def)
-  hence "summable (\<lambda>n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp
-  moreover from r and assms[of "of_real r"] have "\<not>summable (\<lambda>n. f n * of_real r ^ n)" by simp
-  ultimately show False by contradiction
-qed
-
-lemma conv_radius_inftyI':
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "\<And>r. r > c \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
-  shows   "conv_radius f = \<infinity>"
-proof -
-  {
-    fix r :: real
-    have "max r (c + 1) > c" by (auto simp: max_def)
-    from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\<lambda>n. f n * z^n)" by blast
-    from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \<ge> r" by simp
-  }
-  from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \<infinity>"
-    by (cases "conv_radius f") simp_all
-qed
-
-lemma conv_radius_inftyI:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "\<And>r. \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
-  shows   "conv_radius f = \<infinity>"
-  using assms by (rule conv_radius_inftyI')
-
-lemma conv_radius_inftyI'':
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "\<And>z. summable (\<lambda>n. f n * z^n)"
-  shows   "conv_radius f = \<infinity>"
-proof (rule conv_radius_inftyI')
-  fix r :: real assume "r > 0"
-  with assms show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
-    by (intro exI[of _ "of_real r"]) simp
-qed
-
-lemma conv_radius_ratio_limit_ereal:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
-  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
-  shows   "conv_radius f = c"
-proof (rule conv_radius_eqI')
-  show "c \<ge> 0" by (intro Lim_bounded2_ereal[OF lim]) simp_all
-next
-  fix r assume r: "0 < r" "ereal r < c"
-  let ?l = "liminf (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
-  have "?l = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
-    using r by (simp add: norm_mult norm_power divide_simps)
-  also from r have "\<dots> = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
-    by (intro Liminf_ereal_mult_right) simp_all
-  also have "liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
-    by (intro lim_imp_Liminf lim) simp
-  finally have l: "?l = c * ereal (inverse r)" by simp
-  from r have  l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps)
-  show "summable (\<lambda>n. f n * of_real r^n)"
-    by (rule summable_norm_cancel, rule ratio_test_convergence)
-       (insert r nz l l', auto elim!: eventually_mono)
-next
-  fix r assume r: "0 < r" "ereal r > c"
-  let ?l = "limsup (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
-  have "?l = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
-    using r by (simp add: norm_mult norm_power divide_simps)
-  also from r have "\<dots> = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
-    by (intro Limsup_ereal_mult_right) simp_all
-  also have "limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
-    by (intro lim_imp_Limsup lim) simp
-  finally have l: "?l = c * ereal (inverse r)" by simp
-  from r have  l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps)
-  show "\<not>summable (\<lambda>n. norm (f n * of_real r^n))"
-    by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono)
-qed
-
-lemma conv_radius_ratio_limit_ereal_nonzero:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes nz:  "c \<noteq> 0"
-  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
-  shows   "conv_radius f = c"
-proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr)
-  assume "\<not>eventually (\<lambda>n. f n \<noteq> 0) sequentially"
-  hence "frequently (\<lambda>n. f n = 0) sequentially" by (simp add: frequently_def)
-  hence "frequently (\<lambda>n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially"
-    by (force elim!: frequently_elim1)
-  hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto
-  with nz show False by contradiction
-qed 
-
-lemma conv_radius_ratio_limit:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "c' = ereal c"
-  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
-  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
-  shows   "conv_radius f = c'"
-  using assms by (intro conv_radius_ratio_limit_ereal) simp_all
-  
-lemma conv_radius_ratio_limit_nonzero:
-  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
-  assumes "c' = ereal c"
-  assumes nz:  "c \<noteq> 0"
-  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
-  shows   "conv_radius f = c'"
-  using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all
-
-lemma conv_radius_mult_power: 
-  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
-  shows   "conv_radius (\<lambda>n. c ^ n * f n) = conv_radius f / ereal (norm c)"
-proof - 
-  have "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
-          limsup (\<lambda>n. ereal (norm c) * ereal (root n (norm (f n))))" 
-    using eventually_gt_at_top[of "0::nat"]
-    by (intro Limsup_eq) 
-       (auto elim!: eventually_mono simp: norm_mult norm_power real_root_mult real_root_power)
-  also have "\<dots> = ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))"
-    using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all
-  finally have A: "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) = 
-                       ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))" .
-  show ?thesis using assms
-    apply (cases "limsup (\<lambda>n. ereal (root n (norm (f n)))) = 0")
-    apply (simp add: A conv_radius_def)
-    apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult)
-    done
-qed
-
-lemma conv_radius_mult_power_right: 
-  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
-  shows   "conv_radius (\<lambda>n. f n * c ^ n) = conv_radius f / ereal (norm c)"
-  using conv_radius_mult_power[OF assms, of f]
-  unfolding conv_radius_def by (simp add: mult.commute norm_mult)
-
-lemma conv_radius_divide_power: 
-  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
-  shows   "conv_radius (\<lambda>n. f n / c^n) = conv_radius f * ereal (norm c)"
-proof - 
-  from assms have "inverse c \<noteq> 0" by simp
-  from conv_radius_mult_power_right[OF this, of f] show ?thesis
-    by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse)
-qed
-
-
-lemma conv_radius_add_ge: 
-  "min (conv_radius f) (conv_radius g) \<le> 
-       conv_radius (\<lambda>x. f x + g x :: 'a :: {banach,real_normed_div_algebra})"
-  by (rule conv_radius_geI_ex')
-     (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius)
-
-lemma conv_radius_mult_ge:
-  fixes f g :: "nat \<Rightarrow> ('a :: {banach,real_normed_div_algebra})"
-  shows "conv_radius (\<lambda>x. \<Sum>i\<le>x. f i * g (x - i)) \<ge> min (conv_radius f) (conv_radius g)"
-proof (rule conv_radius_geI_ex')
-  fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)"
-  from r have "summable (\<lambda>n. (\<Sum>i\<le>n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))"
-    by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all
-  thus "summable (\<lambda>n. (\<Sum>i\<le>n. f i * g (n - i)) * of_real r ^ n)"
-    by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_setsum_right)
-qed
-
-lemma le_conv_radius_iff:
-  fixes a :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}"
-  shows "r \<le> conv_radius a \<longleftrightarrow> (\<forall>x. norm (x-\<xi>) < r \<longrightarrow> summable (\<lambda>i. a i * (x - \<xi>) ^ i))"
-apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex)
-apply (meson less_ereal.simps(1) not_le order_trans)
-apply (rule_tac x="of_real ra" in exI, simp)
-apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real)
-done
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Summation_Tests.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,920 @@
+(*  Title:    HOL/Multivariate_Analysis/Summation.thy
+    Author:   Manuel Eberl, TU München
+*)
+
+section \<open>Radius of Convergence and Summation Tests\<close>
+
+theory Summation_Tests
+imports
+  Complex_Main
+  "~~/src/HOL/Library/Extended_Real"
+  "~~/src/HOL/Library/Liminf_Limsup"
+begin
+
+text \<open>
+  The definition of the radius of convergence of a power series,
+  various summability tests, lemmas to compute the radius of convergence etc.
+\<close>
+
+subsection \<open>Rounded dual logarithm\<close>
+
+(* This is required for the Cauchy condensation criterion *)
+
+definition "natlog2 n = (if n = 0 then 0 else nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
+
+lemma natlog2_0 [simp]: "natlog2 0 = 0" by (simp add: natlog2_def)
+lemma natlog2_1 [simp]: "natlog2 1 = 0" by (simp add: natlog2_def)
+lemma natlog2_eq_0_iff: "natlog2 n = 0 \<longleftrightarrow> n < 2" by (simp add: natlog2_def)
+
+lemma natlog2_power_of_two [simp]: "natlog2 (2 ^ n) = n"
+  by (simp add: natlog2_def log_nat_power)
+
+lemma natlog2_mono: "m \<le> n \<Longrightarrow> natlog2 m \<le> natlog2 n"
+  unfolding natlog2_def by (simp_all add: nat_mono floor_mono)
+
+lemma pow_natlog2_le: "n > 0 \<Longrightarrow> 2 ^ natlog2 n \<le> n"
+proof -
+  assume n: "n > 0"
+  from n have "of_nat (2 ^ natlog2 n) = 2 powr real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
+    by (subst powr_realpow) (simp_all add: natlog2_def)
+  also have "\<dots> = 2 powr of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>" using n by simp
+  also have "\<dots> \<le> 2 powr log 2 (real_of_nat n)" by (intro powr_mono) (linarith, simp_all)
+  also have "\<dots> = of_nat n" using n by simp
+  finally show ?thesis by simp
+qed
+
+lemma pow_natlog2_gt: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n > n"
+  and pow_natlog2_ge: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n \<ge> n"
+proof -
+  assume n: "n > 0"
+  from n have "of_nat n = 2 powr log 2 (real_of_nat n)" by simp
+  also have "\<dots> < 2 powr (1 + of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
+    by (intro powr_less_mono) (linarith, simp_all)
+  also from n have "\<dots> = 2 powr (1 + real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>))" by simp
+  also from n have "\<dots> = of_nat (2 * 2 ^ natlog2 n)"
+    by (simp_all add: natlog2_def powr_real_of_int powr_add)
+  finally show "2 * 2 ^ natlog2 n > n" by (rule of_nat_less_imp_less)
+  thus "2 * 2 ^ natlog2 n \<ge> n" by simp
+qed
+
+lemma natlog2_eqI:
+  assumes "n > 0" "2^k \<le> n" "n < 2 * 2^k"
+  shows   "natlog2 n = k"
+proof -
+  from assms have "of_nat (2 ^ k) \<le> real_of_nat n"  by (subst of_nat_le_iff) simp_all
+  hence "real_of_int (int k) \<le> log (of_nat 2) (real_of_nat n)"
+    by (subst le_log_iff) (simp_all add: powr_realpow assms del: of_nat_le_iff)
+  moreover from assms have "real_of_nat n < of_nat (2 ^ Suc k)" by (subst of_nat_less_iff) simp_all
+  hence "log 2 (real_of_nat n) < of_nat k + 1"
+    by (subst log_less_iff) (simp_all add: assms powr_realpow powr_add)
+  ultimately have "\<lfloor>log 2 (real_of_nat n)\<rfloor> = of_nat k" by (intro floor_unique) simp_all
+  with assms show ?thesis by (simp add: natlog2_def)
+qed
+
+lemma natlog2_rec:
+  assumes "n \<ge> 2"
+  shows   "natlog2 n = 1 + natlog2 (n div 2)"
+proof (rule natlog2_eqI)
+  from assms have "2 ^ (1 + natlog2 (n div 2)) \<le> 2 * (n div 2)"
+    by (simp add: pow_natlog2_le)
+  also have "\<dots> \<le> n" by simp
+  finally show "2 ^ (1 + natlog2 (n div 2)) \<le> n" .
+next
+  from assms have "n < 2 * (n div 2 + 1)" by simp
+  also from assms have "(n div 2) < 2 ^ (1 + natlog2 (n div 2))"
+    by (simp add: pow_natlog2_gt)
+  hence "2 * (n div 2 + 1) \<le> 2 * (2 ^ (1 + natlog2 (n div 2)))"
+    by (intro mult_left_mono) simp_all
+  finally show "n < 2 * 2 ^ (1 + natlog2 (n div 2))" .
+qed (insert assms, simp_all)
+
+fun natlog2_aux where
+  "natlog2_aux n acc = (if (n::nat) < 2 then acc else natlog2_aux (n div 2) (acc + 1))"
+
+lemma natlog2_aux_correct:
+  "natlog2_aux n acc = acc + natlog2 n"
+  by (induction n acc rule: natlog2_aux.induct) (auto simp: natlog2_rec natlog2_eq_0_iff)
+
+lemma natlog2_code [code]: "natlog2 n = natlog2_aux n 0"
+  by (subst natlog2_aux_correct) simp
+
+
+subsection \<open>Convergence tests for infinite sums\<close>
+
+subsubsection \<open>Root test\<close>
+
+lemma limsup_root_powser:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  shows "limsup (\<lambda>n. ereal (root n (norm (f n * z ^ n)))) =
+             limsup (\<lambda>n. ereal (root n (norm (f n)))) * ereal (norm z)"
+proof -
+  have A: "(\<lambda>n. ereal (root n (norm (f n * z ^ n)))) =
+              (\<lambda>n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h")
+  proof
+    fix n show "?g n = ?h n"
+    by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power)
+  qed
+  show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all
+qed
+
+lemma limsup_root_limit:
+  assumes "(\<lambda>n. ereal (root n (norm (f n)))) \<longlonglongrightarrow> l" (is "?g \<longlonglongrightarrow> _")
+  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = l"
+proof -
+  from assms have "convergent ?g" "lim ?g = l"
+    unfolding convergent_def by (blast intro: limI)+
+  with convergent_limsup_cl show ?thesis by force
+qed
+
+lemma limsup_root_limit':
+  assumes "(\<lambda>n. root n (norm (f n))) \<longlonglongrightarrow> l"
+  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = ereal l"
+  by (intro limsup_root_limit tendsto_ereal assms)
+
+lemma root_test_convergence':
+  fixes f :: "nat \<Rightarrow> 'a :: banach"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  assumes l: "l < 1"
+  shows   "summable f"
+proof -
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have "l \<ge> 0" by simp
+  with l obtain l' where l': "l = ereal l'" by (cases l) simp_all
+
+  define c where "c = (1 - l') / 2"
+  from l and \<open>l \<ge> 0\<close> have c: "l + c > l" "l' + c \<ge> 0" "l' + c < 1" unfolding c_def
+    by (simp_all add: field_simps l')
+  have "\<forall>C>l. eventually (\<lambda>n. ereal (root n (norm (f n))) < C) sequentially"
+    by (subst Limsup_le_iff[symmetric]) (simp add: l_def)
+  with c have "eventually (\<lambda>n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp
+  with eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. norm (f n) \<le> (l' + c) ^ n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    assume "ereal (root n (norm (f n))) < l + ereal c"
+    hence "root n (norm (f n)) \<le> l' + c" by (simp add: l')
+    with c n have "root n (norm (f n)) ^ n \<le> (l' + c) ^ n"
+      by (intro power_mono) (simp_all add: real_root_ge_zero)
+    also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp
+    finally show "norm (f n) \<le> (l' + c) ^ n" by simp
+  qed
+  thus ?thesis
+    by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c)
+qed
+
+lemma root_test_divergence:
+  fixes f :: "nat \<Rightarrow> 'a :: banach"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  assumes l: "l > 1"
+  shows   "\<not>summable f"
+proof
+  assume "summable f"
+  hence bounded: "Bseq f" by (simp add: summable_imp_Bseq)
+
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" by simp
+
+  define c where "c = (if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2)"
+  from l l_nonneg consider "l = \<infinity>" | "\<exists>l'. l = ereal l'" by (cases l) simp_all
+  hence c: "c > 1 \<and> ereal c < l" by cases (insert l, auto simp: c_def field_simps)
+
+  have unbounded: "\<not>bdd_above {n. root n (norm (f n)) > c}"
+  proof
+    assume "bdd_above {n. root n (norm (f n)) > c}"
+    then obtain N where "\<forall>n. root n (norm (f n)) > c \<longrightarrow> n \<le> N" unfolding bdd_above_def by blast
+    hence "\<exists>N. \<forall>n\<ge>N. root n (norm (f n)) \<le> c"
+      by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric])
+    hence "eventually (\<lambda>n. root n (norm (f n)) \<le> c) sequentially"
+      by (auto simp: eventually_at_top_linorder)
+    hence "l \<le> c" unfolding l_def by (intro Limsup_bounded) simp_all
+    with c show False by auto
+  qed
+
+  from bounded obtain K where K: "K > 0" "\<And>n. norm (f n) \<le> K" using BseqE by blast
+  define n where "n = nat \<lceil>log c K\<rceil>"
+  from unbounded have "\<exists>m>n. c < root m (norm (f m))" unfolding bdd_above_def
+    by (auto simp: not_le)
+  then guess m by (elim exE conjE) note m = this
+  from c K have "K = c powr log c K" by (simp add: powr_def log_def)
+  also from c have "c powr log c K \<le> c powr real n" unfolding n_def
+    by (intro powr_mono, linarith, simp)
+  finally have "K \<le> c ^ n" using c by (simp add: powr_realpow)
+  also from c m have "c ^ n < c ^ m" by simp
+  also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all
+  also from m have "... = norm (f m)" by simp
+  finally show False using K(2)[of m]  by simp
+qed
+
+
+subsubsection \<open>Cauchy's condensation test\<close>
+
+context
+fixes f :: "nat \<Rightarrow> real"
+begin
+
+private lemma condensation_inequality:
+  assumes mono: "\<And>m n. 0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> f n \<le> f m"
+  shows   "(\<Sum>k=1..<n. f k) \<ge> (\<Sum>k=1..<n. f (2 * 2 ^ natlog2 k))" (is "?thesis1")
+          "(\<Sum>k=1..<n. f k) \<le> (\<Sum>k=1..<n. f (2 ^ natlog2 k))" (is "?thesis2")
+  by (intro setsum_mono mono pow_natlog2_ge pow_natlog2_le, simp, simp)+
+
+private lemma condensation_condense1: "(\<Sum>k=1..<2^n. f (2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ k))"
+proof (induction n)
+  case (Suc n)
+  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
+  also have "(\<Sum>k\<in>\<dots>. f (2 ^ natlog2 k)) =
+                 (\<Sum>k<n. 2^k * f (2^k)) + (\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k))"
+    by (subst setsum.union_disjoint) (insert Suc, auto)
+  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
+  hence "(\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
+    by (intro setsum.cong) simp_all
+  also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
+  finally show ?case by simp
+qed simp
+
+private lemma condensation_condense2: "(\<Sum>k=1..<2^n. f (2 * 2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ Suc k))"
+proof (induction n)
+  case (Suc n)
+  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto
+  also have "(\<Sum>k\<in>\<dots>. f (2 * 2 ^ natlog2 k)) =
+                 (\<Sum>k<n. 2^k * f (2^Suc k)) + (\<Sum>k = 2^n..<2^Suc n. f (2 * 2^natlog2 k))"
+    by (subst setsum.union_disjoint) (insert Suc, auto)
+  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
+  hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
+    by (intro setsum.cong) simp_all
+  also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
+  finally show ?case by simp
+qed simp
+
+lemma condensation_test:
+  assumes mono: "\<And>m. 0 < m \<Longrightarrow> f (Suc m) \<le> f m"
+  assumes nonneg: "\<And>n. f n \<ge> 0"
+  shows "summable f \<longleftrightarrow> summable (\<lambda>n. 2^n * f (2^n))"
+proof -
+  define f' where "f' n = (if n = 0 then 0 else f n)" for n
+  from mono have mono': "decseq (\<lambda>n. f (Suc n))" by (intro decseq_SucI) simp
+  hence mono': "f n \<le> f m" if "m \<le> n" "m > 0" for m n
+    using that decseqD[OF mono', of "m - 1" "n - 1"] by simp
+
+  have "(\<lambda>n. f (Suc n)) = (\<lambda>n. f' (Suc n))" by (intro ext) (simp add: f'_def)
+  hence "summable f \<longleftrightarrow> summable f'"
+    by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:)
+  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k<n. f' k)" unfolding summable_iff_convergent ..
+  also have "monoseq (\<lambda>n. \<Sum>k<n. f' k)" unfolding f'_def
+    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
+  hence "convergent (\<lambda>n. \<Sum>k<n. f' k) \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. f' k)"
+    by (rule monoseq_imp_convergent_iff_Bseq)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f' k)" unfolding One_nat_def
+    by (subst setsum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f k)" unfolding f'_def by simp
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
+    by (rule nonneg_incseq_Bseq_subseq_iff[symmetric])
+       (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^k))"
+  proof (intro iffI)
+    assume A: "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
+    have "eventually (\<lambda>n. norm (\<Sum>k<n. 2^k * f (2^Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      have "norm (\<Sum>k<n. 2^k * f (2^Suc k)) = (\<Sum>k<n. 2^k * f (2^Suc k))" unfolding real_norm_def
+        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+      also have "\<dots> \<le> (\<Sum>k=1..<2^n. f k)"
+        by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono')
+      also have "\<dots> = norm \<dots>" unfolding real_norm_def
+        by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+      finally show "norm (\<Sum>k<n. 2 ^ k * f (2 ^ Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)" .
+    qed
+    from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
+    from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
+      by (simp add: setsum_right_distrib setsum_left_distrib mult_ac)
+    hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
+      by (intro Bseq_add, subst setsum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
+    hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
+      by (subst setsum_head_upt_Suc) (simp_all add: add_ac)
+    thus "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+      by (subst (asm) Bseq_Suc_iff) (simp add: atLeast0LessThan)
+  next
+    assume A: "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    have "eventually (\<lambda>n. norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      have "norm (\<Sum>k=1..<2^n. f k) = (\<Sum>k=1..<2^n. f k)" unfolding real_norm_def
+        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+      also have "\<dots> \<le> (\<Sum>k<n. 2^k * f (2^k))"
+        by (subst condensation_condense1 [symmetric]) (intro condensation_inequality mono')
+      also have "\<dots> = norm \<dots>" unfolding real_norm_def
+        by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+      finally show "norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))" .
+    qed
+    from this and A show "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)" by (rule Bseq_eventually_mono)
+  qed
+  also have "monoseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
+  hence "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k))) \<longleftrightarrow> convergent (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    by (rule monoseq_imp_convergent_iff_Bseq [symmetric])
+  also have "\<dots> \<longleftrightarrow> summable (\<lambda>k. 2^k * f (2^k))" by (simp only: summable_iff_convergent)
+  finally show ?thesis .
+qed
+
+end
+
+
+subsubsection \<open>Summability of powers\<close>
+
+lemma abs_summable_complex_powr_iff:
+    "summable (\<lambda>n. norm (exp (of_real (ln (of_nat n)) * s))) \<longleftrightarrow> Re s < -1"
+proof (cases "Re s \<le> 0")
+  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
+  case False
+  with eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. norm (1 :: real) \<le> norm (exp (?l n * s))) sequentially"
+    by (auto intro!: ge_one_powr_ge_zero elim!: eventually_mono)
+  from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff)
+next
+  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
+  case True
+  hence "summable (\<lambda>n. norm (exp (?l n * s))) \<longleftrightarrow> summable (\<lambda>n. 2^n * norm (exp (?l (2^n) * s)))"
+    by (intro condensation_test) (auto intro!: mult_right_mono_neg)
+  also have "(\<lambda>n. 2^n * norm (exp (?l (2^n) * s))) = (\<lambda>n. (2 powr (Re s + 1)) ^ n)"
+  proof
+    fix n :: nat
+    have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)"
+      using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps)
+    also have "\<dots> = exp (real n * (ln 2 * (Re s + 1)))"
+      by (simp add: algebra_simps exp_add)
+    also have "\<dots> = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp
+    also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def)
+    finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" .
+  qed
+  also have "summable \<dots> \<longleftrightarrow> 2 powr (Re s + 1) < 2 powr 0"
+    by (subst summable_geometric_iff) simp
+  also have "\<dots> \<longleftrightarrow> Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith)
+  finally show ?thesis .
+qed
+
+lemma summable_complex_powr_iff:
+  assumes "Re s < -1"
+  shows   "summable (\<lambda>n. exp (of_real (ln (of_nat n)) * s))"
+  by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact
+
+lemma summable_real_powr_iff: "summable (\<lambda>n. of_nat n powr s :: real) \<longleftrightarrow> s < -1"
+proof -
+  from eventually_gt_at_top[of "0::nat"]
+    have "summable (\<lambda>n. of_nat n powr s) \<longleftrightarrow> summable (\<lambda>n. exp (ln (of_nat n) * s))"
+    by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def)
+  also have "\<dots> \<longleftrightarrow> s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp
+  finally show ?thesis .
+qed
+
+lemma inverse_power_summable:
+  assumes s: "s \<ge> 2"
+  shows "summable (\<lambda>n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))"
+proof (rule summable_norm_cancel, subst summable_cong)
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top"
+    by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow)
+qed (insert s summable_real_powr_iff[of "-s"], simp_all)
+
+lemma not_summable_harmonic: "\<not>summable (\<lambda>n. inverse (of_nat n) :: 'a :: real_normed_field)"
+proof
+  assume "summable (\<lambda>n. inverse (of_nat n) :: 'a)"
+  hence "convergent (\<lambda>n. norm (of_real (\<Sum>k<n. inverse (of_nat k)) :: 'a))"
+    by (simp add: summable_iff_convergent convergent_norm)
+  hence "convergent (\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real)" by (simp only: norm_of_real)
+  also have "(\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real) = (\<lambda>n. \<Sum>k<n. inverse (of_nat k))"
+    by (intro ext abs_of_nonneg setsum_nonneg) auto
+  also have "convergent \<dots> \<longleftrightarrow> summable (\<lambda>k. inverse (of_nat k) :: real)"
+    by (simp add: summable_iff_convergent)
+  finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus)
+qed
+
+
+subsubsection \<open>Kummer's test\<close>
+
+lemma kummers_test_convergence:
+  fixes f p :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
+  assumes nonneg_p: "eventually (\<lambda>n. p n \<ge> 0) sequentially"
+  defines "l \<equiv> liminf (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
+  assumes l: "l > 0"
+  shows   "summable f"
+  unfolding summable_iff_convergent'
+proof -
+  define r where "r = (if l = \<infinity> then 1 else real_of_ereal l / 2)"
+  from l have "r > 0 \<and> of_real r < l" by (cases l) (simp_all add: r_def)
+  hence r: "r > 0" "of_real r < l" by simp_all
+  hence "eventually (\<lambda>n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially"
+    unfolding l_def by (force dest: less_LiminfD)
+  moreover from pos_f have "eventually (\<lambda>n. f (Suc n) > 0) sequentially"
+    by (subst eventually_sequentially_Suc)
+  ultimately have "eventually (\<lambda>n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially"
+    by eventually_elim (simp add: field_simps)
+  from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]]
+    obtain m where m: "\<And>n. n \<ge> m \<Longrightarrow> f n > 0" "\<And>n. n \<ge> m \<Longrightarrow> p n \<ge> 0"
+        "\<And>n. n \<ge> m \<Longrightarrow> p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)"
+    unfolding eventually_at_top_linorder by blast
+
+  let ?c = "(norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r"
+  have "Bseq (\<lambda>n. (\<Sum>k\<le>n + Suc m. f k))"
+  proof (rule BseqI')
+    fix k :: nat
+    define n where "n = k + Suc m"
+    have n: "n > m" by (simp add: n_def)
+
+    from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
+      by (simp add: setsum_right_distrib[symmetric] abs_mult)
+    also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
+    hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
+    also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
+      by (subst setsum.union_disjoint) auto
+    also have "norm \<dots> \<le> norm (\<Sum>k\<le>m. r * f k) + norm (\<Sum>k=Suc m..n. r * f k)"
+      by (rule norm_triangle_ineq)
+    also from r less_imp_le[OF m(1)] have "(\<Sum>k=Suc m..n. r * f k) \<ge> 0"
+      by (intro setsum_nonneg) auto
+    hence "norm (\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=Suc m..n. r * f k)" by simp
+    also have "(\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=m..<n. r * f (Suc k))"
+     by (subst setsum_shift_bounds_Suc_ivl [symmetric])
+          (simp only: atLeastLessThanSuc_atLeastAtMost)
+    also from m have "\<dots> \<le> (\<Sum>k=m..<n. p k * f k - p (Suc k) * f (Suc k))"
+      by (intro setsum_mono[OF less_imp_le]) simp_all
+    also have "\<dots> = -(\<Sum>k=m..<n. p (Suc k) * f (Suc k) - p k * f k)"
+      by (simp add: setsum_negf [symmetric] algebra_simps)
+    also from n have "\<dots> = p m * f m - p n * f n"
+      by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all
+    also from less_imp_le[OF m(1)] m(2) n have "\<dots> \<le> p m * f m" by simp
+    finally show "norm (\<Sum>k\<le>n. f k) \<le> (norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r" using r
+      by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac)
+  qed
+  moreover have "(\<Sum>k\<le>n. f k) \<le> (\<Sum>k\<le>n'. f k)" if "Suc m \<le> n" "n \<le> n'" for n n'
+    using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto
+  ultimately show "convergent (\<lambda>n. \<Sum>k\<le>n. f k)" by (rule Bseq_monoseq_convergent'_inc)
+qed
+
+
+lemma kummers_test_divergence:
+  fixes f p :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
+  assumes pos_p: "eventually (\<lambda>n. p n > 0) sequentially"
+  assumes divergent_p: "\<not>summable (\<lambda>n. inverse (p n))"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
+  assumes l: "l < 0"
+  shows   "\<not>summable f"
+proof
+  assume "summable f"
+  from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]]
+    obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> p n > 0" "\<And>n. n \<ge> N \<Longrightarrow> f n > 0"
+                      "\<And>n. n \<ge> N \<Longrightarrow> p n * f n / f (Suc n) - p (Suc n) < 0"
+    by (auto simp: eventually_at_top_linorder)
+  hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \<ge> N" for n using that N[of n] N[of "Suc n"]
+    by (simp add: field_simps)
+  have "p n * f n \<ge> p N * f N" if "n \<ge> N" for n using that and A
+      by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans)
+  from eventually_ge_at_top[of N] N this
+    have "eventually (\<lambda>n. norm (p N * f N * inverse (p n)) \<le> f n) sequentially"
+    by (auto elim!: eventually_mono simp: field_simps abs_of_pos)
+  from this and \<open>summable f\<close> have "summable (\<lambda>n. p N * f N * inverse (p n))"
+    by (rule summable_comparison_test_ev)
+  from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl]
+    have "summable (\<lambda>n. inverse (p n))" by (simp add: divide_simps)
+  with divergent_p show False by contradiction
+qed
+
+
+subsubsection \<open>Ratio test\<close>
+
+lemma ratio_test_convergence:
+  fixes f :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> liminf (\<lambda>n. ereal (f n / f (Suc n)))"
+  assumes l: "l > 1"
+  shows   "summable f"
+proof (rule kummers_test_convergence[OF pos_f])
+  note l
+  also have "l = liminf (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1"
+    by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
+  finally show "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) > 0"
+    by (cases "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
+qed simp
+
+lemma ratio_test_divergence:
+  fixes f :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (f n / f (Suc n)))"
+  assumes l: "l < 1"
+  shows   "\<not>summable f"
+proof (rule kummers_test_divergence[OF pos_f])
+  have "limsup (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1 = l"
+    by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
+  also note l
+  finally show "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) < 0"
+    by (cases "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
+qed (simp_all add: summable_const_iff)
+
+
+subsubsection \<open>Raabe's test\<close>
+
+lemma raabes_test_convergence:
+fixes f :: "nat \<Rightarrow> real"
+  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> liminf (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
+  assumes l: "l > 1"
+  shows   "summable f"
+proof (rule kummers_test_convergence)
+  let ?l' = "liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
+  have "1 < l" by fact
+  also have "l = liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
+    by (simp add: l_def algebra_simps)
+  also have "\<dots> = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all
+  finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps)
+qed (simp_all add: pos)
+
+lemma raabes_test_divergence:
+fixes f :: "nat \<Rightarrow> real"
+  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
+  assumes l: "l < 1"
+  shows   "\<not>summable f"
+proof (rule kummers_test_divergence)
+  let ?l' = "limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
+  note l
+  also have "l = limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
+    by (simp add: l_def algebra_simps)
+  also have "\<dots> = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all
+  finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps)
+qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all)
+
+
+
+subsection \<open>Radius of convergence\<close>
+
+text \<open>
+  The radius of convergence of a power series. This value always exists, ranges from
+  @{term "0::ereal"} to @{term "\<infinity>::ereal"}, and the power series is guaranteed to converge for
+  all inputs with a norm that is smaller than that radius and to diverge for all inputs with a
+  norm that is greater.
+\<close>
+definition conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
+  "conv_radius f = inverse (limsup (\<lambda>n. ereal (root n (norm (f n)))))"
+
+lemma conv_radius_nonneg: "conv_radius f \<ge> 0"
+proof -
+  have "0 = limsup (\<lambda>n. 0)" by (subst Limsup_const) simp_all
+  also have "\<dots> \<le> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+    by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally show ?thesis
+    unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff)
+qed
+
+lemma conv_radius_zero [simp]: "conv_radius (\<lambda>_. 0) = \<infinity>"
+  by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const)
+
+lemma conv_radius_cong:
+  assumes "eventually (\<lambda>x. f x = g x) sequentially"
+  shows   "conv_radius f = conv_radius g"
+proof -
+  have "eventually (\<lambda>n. ereal (root n (norm (f n))) = ereal (root n (norm (g n)))) sequentially"
+    using assms by eventually_elim simp
+  from Limsup_eq[OF this] show ?thesis unfolding conv_radius_def by simp
+qed
+
+lemma conv_radius_altdef:
+  "conv_radius f = liminf (\<lambda>n. inverse (ereal (root n (norm (f n)))))"
+  by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def)
+
+lemma abs_summable_in_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) < conv_radius f"
+  shows   "summable (\<lambda>n. norm (f n * z ^ n))"
+proof (rule root_test_convergence')
+  define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" .
+
+  have "limsup (\<lambda>n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def
+    by (rule limsup_root_powser)
+  also from l_nonneg consider "l = 0" | "l = \<infinity>" | "\<exists>l'. l = ereal l' \<and> l' > 0"
+    by (cases "l") (auto simp: less_le)
+  hence "l * ereal (norm z) < 1"
+  proof cases
+    assume "l = \<infinity>"
+    hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp
+    with assms show ?thesis by simp
+  next
+    assume "\<exists>l'. l = ereal l' \<and> l' > 0"
+    then guess l' by (elim exE conjE) note l' = this
+    hence "l \<noteq> \<infinity>" by auto
+    have "l * ereal (norm z) < l * conv_radius f"
+      by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms)
+    also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def)
+    also from l' have "l * inverse l = 1" by simp
+    finally show ?thesis .
+  qed simp_all
+  finally show "limsup (\<lambda>n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp
+qed
+
+lemma summable_in_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) < conv_radius f"
+  shows   "summable (\<lambda>n. f n * z ^ n)"
+  by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+
+
+lemma not_summable_outside_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) > conv_radius f"
+  shows   "\<not>summable (\<lambda>n. f n * z ^ n)"
+proof (rule root_test_divergence)
+  define l where "l = limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" .
+  from assms have l_nz: "l \<noteq> 0" unfolding conv_radius_def l_def by auto
+
+  have "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)"
+    unfolding l_def by (rule limsup_root_powser)
+  also have "... > 1"
+  proof (cases l)
+    assume "l = \<infinity>"
+    with assms conv_radius_nonneg[of f] show ?thesis
+      by (auto simp: zero_ereal_def[symmetric])
+  next
+    fix l' assume l': "l = ereal l'"
+    from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps)
+    also from l_nz have "inverse l = conv_radius f"
+      unfolding l_def conv_radius_def by auto
+    also from l' l_nz l_nonneg assms have "l * \<dots> < l * ereal (norm z)"
+      by (intro ereal_mult_strict_left_mono) (auto simp: l')
+    finally show ?thesis .
+  qed (insert l_nonneg, simp_all)
+  finally show "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) > 1" .
+qed
+
+
+lemma conv_radius_geI:
+  assumes "summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius f \<ge> norm z"
+  using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_leI:
+  assumes "\<not>summable (\<lambda>n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))"
+  shows   "conv_radius f \<le> norm z"
+  using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_leI':
+  assumes "\<not>summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius f \<le> norm z"
+  using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_geI_ex:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f \<ge> R"
+proof (rule linorder_cases[of "conv_radius f" R])
+  assume R: "conv_radius f < R"
+  with conv_radius_nonneg[of f] obtain conv_radius'
+    where [simp]: "conv_radius f = ereal conv_radius'"
+    by (cases "conv_radius f") simp_all
+  define r where "r = (if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)"
+  from R conv_radius_nonneg[of f] have "0 < r \<and> ereal r < R \<and> ereal r > conv_radius f"
+    unfolding r_def by (cases R) (auto simp: r_def field_simps)
+  with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\<lambda>n. f n * z^n)" by auto
+  with not_summable_outside_conv_radius[of f z] show ?thesis by simp
+qed simp_all
+
+lemma conv_radius_geI_ex':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * of_real r^n)"
+  shows   "conv_radius f \<ge> R"
+proof (rule conv_radius_geI_ex)
+  fix r assume "0 < r" "ereal r < R"
+  with assms[of r] show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)"
+    by (intro exI[of _ "of_real r :: 'a"]) auto
+qed
+
+lemma conv_radius_leI_ex:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
+  shows   "conv_radius f \<le> R"
+proof (rule linorder_cases[of "conv_radius f" R])
+  assume R: "conv_radius f > R"
+  from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all
+  define r where
+    "r = (if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)"
+  from R conv_radius_nonneg[of f] have "r > R \<and> r < conv_radius f" unfolding r_def
+    by (cases "conv_radius f") (auto simp: r_def field_simps R')
+  with assms(1) assms(2)[of r] R'
+    obtain z where "norm z < conv_radius f" "\<not>summable (\<lambda>n. norm (f n * z^n))" by auto
+  with abs_summable_in_conv_radius[of z f] show ?thesis by auto
+qed simp_all
+
+lemma conv_radius_leI_ex':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. f n * of_real r^n)"
+  shows   "conv_radius f \<le> R"
+proof (rule conv_radius_leI_ex)
+  fix r assume "0 < r" "ereal r > R"
+  with assms(2)[of r] show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))"
+    by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel)
+qed fact+
+
+lemma conv_radius_eqI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
+  shows   "conv_radius f = R"
+  by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms)
+
+lemma conv_radius_eqI':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * (of_real r)^n)"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. norm (f n * (of_real r)^n))"
+  shows   "conv_radius f = R"
+proof (intro conv_radius_eqI[OF assms(1)])
+  fix r assume "0 < r" "ereal r < R" with assms(2)[OF this]
+    show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)" by force
+next
+  fix r assume "0 < r" "ereal r > R" with assms(3)[OF this]
+    show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))" by force
+qed
+
+lemma conv_radius_zeroI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>z. z \<noteq> 0 \<Longrightarrow> \<not>summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = 0"
+proof (rule ccontr)
+  assume "conv_radius f \<noteq> 0"
+  with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp
+  define r where "r = (if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2)"
+  from pos have r: "ereal r > 0 \<and> ereal r < conv_radius f"
+    by (cases "conv_radius f") (simp_all add: r_def)
+  hence "summable (\<lambda>n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp
+  moreover from r and assms[of "of_real r"] have "\<not>summable (\<lambda>n. f n * of_real r ^ n)" by simp
+  ultimately show False by contradiction
+qed
+
+lemma conv_radius_inftyI':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>r. r > c \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+proof -
+  {
+    fix r :: real
+    have "max r (c + 1) > c" by (auto simp: max_def)
+    from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\<lambda>n. f n * z^n)" by blast
+    from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \<ge> r" by simp
+  }
+  from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \<infinity>"
+    by (cases "conv_radius f") simp_all
+qed
+
+lemma conv_radius_inftyI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>r. \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+  using assms by (rule conv_radius_inftyI')
+
+lemma conv_radius_inftyI'':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>z. summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+proof (rule conv_radius_inftyI')
+  fix r :: real assume "r > 0"
+  with assms show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+    by (intro exI[of _ "of_real r"]) simp
+qed
+
+lemma conv_radius_ratio_limit_ereal:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c"
+proof (rule conv_radius_eqI')
+  show "c \<ge> 0" by (intro Lim_bounded2_ereal[OF lim]) simp_all
+next
+  fix r assume r: "0 < r" "ereal r < c"
+  let ?l = "liminf (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
+  have "?l = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
+    using r by (simp add: norm_mult norm_power divide_simps)
+  also from r have "\<dots> = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
+    by (intro Liminf_ereal_mult_right) simp_all
+  also have "liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
+    by (intro lim_imp_Liminf lim) simp
+  finally have l: "?l = c * ereal (inverse r)" by simp
+  from r have  l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps)
+  show "summable (\<lambda>n. f n * of_real r^n)"
+    by (rule summable_norm_cancel, rule ratio_test_convergence)
+       (insert r nz l l', auto elim!: eventually_mono)
+next
+  fix r assume r: "0 < r" "ereal r > c"
+  let ?l = "limsup (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
+  have "?l = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
+    using r by (simp add: norm_mult norm_power divide_simps)
+  also from r have "\<dots> = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
+    by (intro Limsup_ereal_mult_right) simp_all
+  also have "limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
+    by (intro lim_imp_Limsup lim) simp
+  finally have l: "?l = c * ereal (inverse r)" by simp
+  from r have  l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps)
+  show "\<not>summable (\<lambda>n. norm (f n * of_real r^n))"
+    by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono)
+qed
+
+lemma conv_radius_ratio_limit_ereal_nonzero:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes nz:  "c \<noteq> 0"
+  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c"
+proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr)
+  assume "\<not>eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  hence "frequently (\<lambda>n. f n = 0) sequentially" by (simp add: frequently_def)
+  hence "frequently (\<lambda>n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially"
+    by (force elim!: frequently_elim1)
+  hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto
+  with nz show False by contradiction
+qed
+
+lemma conv_radius_ratio_limit:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "c' = ereal c"
+  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c'"
+  using assms by (intro conv_radius_ratio_limit_ereal) simp_all
+
+lemma conv_radius_ratio_limit_nonzero:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "c' = ereal c"
+  assumes nz:  "c \<noteq> 0"
+  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c'"
+  using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all
+
+lemma conv_radius_mult_power:
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. c ^ n * f n) = conv_radius f / ereal (norm c)"
+proof -
+  have "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
+          limsup (\<lambda>n. ereal (norm c) * ereal (root n (norm (f n))))"
+    using eventually_gt_at_top[of "0::nat"]
+    by (intro Limsup_eq)
+       (auto elim!: eventually_mono simp: norm_mult norm_power real_root_mult real_root_power)
+  also have "\<dots> = ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))"
+    using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all
+  finally have A: "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
+                       ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))" .
+  show ?thesis using assms
+    apply (cases "limsup (\<lambda>n. ereal (root n (norm (f n)))) = 0")
+    apply (simp add: A conv_radius_def)
+    apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult)
+    done
+qed
+
+lemma conv_radius_mult_power_right:
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. f n * c ^ n) = conv_radius f / ereal (norm c)"
+  using conv_radius_mult_power[OF assms, of f]
+  unfolding conv_radius_def by (simp add: mult.commute norm_mult)
+
+lemma conv_radius_divide_power:
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. f n / c^n) = conv_radius f * ereal (norm c)"
+proof -
+  from assms have "inverse c \<noteq> 0" by simp
+  from conv_radius_mult_power_right[OF this, of f] show ?thesis
+    by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse)
+qed
+
+
+lemma conv_radius_add_ge:
+  "min (conv_radius f) (conv_radius g) \<le>
+       conv_radius (\<lambda>x. f x + g x :: 'a :: {banach,real_normed_div_algebra})"
+  by (rule conv_radius_geI_ex')
+     (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius)
+
+lemma conv_radius_mult_ge:
+  fixes f g :: "nat \<Rightarrow> ('a :: {banach,real_normed_div_algebra})"
+  shows "conv_radius (\<lambda>x. \<Sum>i\<le>x. f i * g (x - i)) \<ge> min (conv_radius f) (conv_radius g)"
+proof (rule conv_radius_geI_ex')
+  fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)"
+  from r have "summable (\<lambda>n. (\<Sum>i\<le>n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))"
+    by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all
+  thus "summable (\<lambda>n. (\<Sum>i\<le>n. f i * g (n - i)) * of_real r ^ n)"
+    by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_setsum_right)
+qed
+
+lemma le_conv_radius_iff:
+  fixes a :: "nat \<Rightarrow> 'a::{real_normed_div_algebra,banach}"
+  shows "r \<le> conv_radius a \<longleftrightarrow> (\<forall>x. norm (x-\<xi>) < r \<longrightarrow> summable (\<lambda>i. a i * (x - \<xi>) ^ i))"
+apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex)
+apply (meson less_ereal.simps(1) not_le order_trans)
+apply (rule_tac x="of_real ra" in exI, simp)
+apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real)
+done
+
+end
--- a/src/HOL/Multivariate_Analysis/Uniform_Limit.thy	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Uniform_Limit.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -6,7 +6,7 @@
 section \<open>Uniform Limit and Uniform Convergence\<close>
 
 theory Uniform_Limit
-imports Topology_Euclidean_Space Summation
+imports Topology_Euclidean_Space Summation_Tests
 begin
 
 definition uniformly_on :: "'a set \<Rightarrow> ('a \<Rightarrow> 'b::metric_space) \<Rightarrow> ('a \<Rightarrow> 'b) filter"
@@ -18,9 +18,9 @@
 definition uniformly_convergent_on where
   "uniformly_convergent_on X f \<longleftrightarrow> (\<exists>l. uniform_limit X f l sequentially)"
 
-definition uniformly_Cauchy_on where 
+definition uniformly_Cauchy_on where
   "uniformly_Cauchy_on X f \<longleftrightarrow> (\<forall>e>0. \<exists>M. \<forall>x\<in>X. \<forall>(m::nat)\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e)"
-  
+
 lemma uniform_limit_iff:
   "uniform_limit S f l F \<longleftrightarrow> (\<forall>e>0. \<forall>\<^sub>F n in F. \<forall>x\<in>S. dist (f n x) (l x) < e)"
   unfolding filterlim_iff uniformly_on_def
@@ -141,7 +141,7 @@
   shows   "uniformly_Cauchy_on X f"
 proof (rule uniformly_Cauchy_onI)
   fix e :: real assume e: "e > 0"
-  from assms[OF this] obtain M 
+  from assms[OF this] obtain M
     where M: "\<And>x m n. x \<in> X \<Longrightarrow> m \<ge> M \<Longrightarrow> n > m \<Longrightarrow> dist (f m x) (f n x) < e" by fast
   {
     fix x m n assume x: "x \<in> X" and m: "m \<ge> M" and n: "n \<ge> M"
@@ -151,7 +151,7 @@
   thus "\<exists>M. \<forall>x\<in>X. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (f m x) (f n x) < e" by fast
 qed
 
-lemma uniformly_Cauchy_imp_Cauchy: 
+lemma uniformly_Cauchy_imp_Cauchy:
   "uniformly_Cauchy_on X f \<Longrightarrow> x \<in> X \<Longrightarrow> Cauchy (\<lambda>n. f n x)"
   unfolding Cauchy_def uniformly_Cauchy_on_def by fast
 
@@ -167,7 +167,7 @@
        and B: "\<And>x. x \<in> X \<Longrightarrow> h x = i x"
     {
       fix e ::real assume "e > 0"
-      with C have "eventually (\<lambda>y. \<forall>x\<in>X. dist (f y x) (h x) < e) F" 
+      with C have "eventually (\<lambda>y. \<forall>x\<in>X. dist (f y x) (h x) < e) F"
         unfolding uniform_limit_iff by blast
       with A have "eventually (\<lambda>y. \<forall>x\<in>X. dist (g y x) (i x) < e) F"
         by eventually_elim (insert B, simp_all)
@@ -188,7 +188,7 @@
   "uniformly_convergent_on X f \<longleftrightarrow> uniform_limit X f (\<lambda>x. lim (\<lambda>n. f n x)) sequentially"
 proof
   assume "uniformly_convergent_on X f"
-  then obtain l where l: "uniform_limit X f l sequentially" 
+  then obtain l where l: "uniform_limit X f l sequentially"
     unfolding uniformly_convergent_on_def by blast
   from l have "uniform_limit X f (\<lambda>x. lim (\<lambda>n. f n x)) sequentially \<longleftrightarrow>
                       uniform_limit X f l sequentially"
@@ -221,11 +221,11 @@
     show "\<forall>x\<in>X. dist (f n x) (?f x) < e"
     proof
       fix x assume x: "x \<in> X"
-      with assms have "(\<lambda>n. f n x) \<longlonglongrightarrow> ?f x" 
+      with assms have "(\<lambda>n. f n x) \<longlonglongrightarrow> ?f x"
         by (auto dest!: Cauchy_convergent uniformly_Cauchy_imp_Cauchy simp: convergent_LIMSEQ_iff)
       with \<open>e/2 > 0\<close> have "eventually (\<lambda>m. m \<ge> N \<and> dist (f m x) (?f x) < e/2) sequentially"
         by (intro tendstoD eventually_conj eventually_ge_at_top)
-      then obtain m where m: "m \<ge> N" "dist (f m x) (?f x) < e/2" 
+      then obtain m where m: "m \<ge> N" "dist (f m x) (?f x) < e/2"
         unfolding eventually_at_top_linorder by blast
       have "dist (f n x) (?f x) \<le> dist (f n x) (f m x) + dist (f m x) (?f x)"
           by (rule dist_triangle)
@@ -306,16 +306,16 @@
   assumes "summable M"
   shows "uniform_limit A (\<lambda>n x. \<Sum>i<n. f i x) (\<lambda>x. suminf (\<lambda>i. f i x)) sequentially"
   using assms by (intro weierstrass_m_test_ev always_eventually) auto
-    
+
 corollary weierstrass_m_test'_ev:
   fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: banach"
-  assumes "eventually (\<lambda>n. \<forall>x\<in>A. norm (f n x) \<le> M n) sequentially" "summable M" 
+  assumes "eventually (\<lambda>n. \<forall>x\<in>A. norm (f n x) \<le> M n) sequentially" "summable M"
   shows   "uniformly_convergent_on A (\<lambda>n x. \<Sum>i<n. f i x)"
   unfolding uniformly_convergent_on_def by (rule exI, rule weierstrass_m_test_ev[OF assms])
 
 corollary weierstrass_m_test':
   fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: banach"
-  assumes "\<And>n x. x \<in> A \<Longrightarrow> norm (f n x) \<le> M n" "summable M" 
+  assumes "\<And>n x. x \<in> A \<Longrightarrow> norm (f n x) \<le> M n" "summable M"
   shows   "uniformly_convergent_on A (\<lambda>n x. \<Sum>i<n. f i x)"
   unfolding uniformly_convergent_on_def by (rule exI, rule weierstrass_m_test[OF assms])
 
@@ -518,7 +518,7 @@
   unfolding uniformly_convergent_on_def by (blast dest: uniform_limit_minus)
 
 lemma uniformly_convergent_mult:
-  "uniformly_convergent_on A f \<Longrightarrow> 
+  "uniformly_convergent_on A f \<Longrightarrow>
       uniformly_convergent_on A (\<lambda>k x. c * f k x :: 'a :: {real_normed_algebra})"
   unfolding uniformly_convergent_on_def
   by (blast dest: bounded_linear_uniform_limit_intros(13))
--- a/src/HOL/Multivariate_Analysis/Weierstrass.thy	Thu Aug 04 18:45:28 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1218 +0,0 @@
-section \<open>The Bernstein-Weierstrass and Stone-Weierstrass Theorems\<close>
-
-text\<open>By L C Paulson (2015)\<close>
-
-theory Weierstrass
-imports Uniform_Limit Path_Connected
-begin
-
-subsection \<open>Bernstein polynomials\<close>
-
-definition Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
-  "Bernstein n k x \<equiv> of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)"
-
-lemma Bernstein_nonneg: "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> Bernstein n k x"
-  by (simp add: Bernstein_def)
-
-lemma Bernstein_pos: "\<lbrakk>0 < x; x < 1; k \<le> n\<rbrakk> \<Longrightarrow> 0 < Bernstein n k x"
-  by (simp add: Bernstein_def)
-
-lemma sum_Bernstein [simp]: "(\<Sum> k = 0..n. Bernstein n k x) = 1"
-  using binomial_ring [of x "1-x" n]
-  by (simp add: Bernstein_def)
-
-lemma binomial_deriv1:
-    "(\<Sum>k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)"
-  apply (rule DERIV_unique [where f = "\<lambda>a. (a+b)^n" and x=a])
-  apply (subst binomial_ring)
-  apply (rule derivative_eq_intros setsum.cong | simp)+
-  done
-
-lemma binomial_deriv2:
-    "(\<Sum>k=0..n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) =
-     of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)"
-  apply (rule DERIV_unique [where f = "\<lambda>a. of_nat n * (a+b::real) ^ (n-1)" and x=a])
-  apply (subst binomial_deriv1 [symmetric])
-  apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+
-  done
-
-lemma sum_k_Bernstein [simp]: "(\<Sum>k = 0..n. real k * Bernstein n k x) = of_nat n * x"
-  apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
-  apply (simp add: setsum_left_distrib)
-  apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong)
-  done
-
-lemma sum_kk_Bernstein [simp]: "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2"
-proof -
-  have "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
-    apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric])
-    apply (simp add: setsum_left_distrib)
-    apply (rule setsum.cong [OF refl])
-    apply (simp add: Bernstein_def power2_eq_square algebra_simps)
-    apply (rename_tac k)
-    apply (subgoal_tac "k = 0 \<or> k = 1 \<or> (\<exists>k'. k = Suc (Suc k'))")
-    apply (force simp add: field_simps of_nat_Suc power2_eq_square)
-    by presburger
-  also have "... = n * (n - 1) * x\<^sup>2"
-    by auto
-  finally show ?thesis
-    by auto
-qed
-
-subsection \<open>Explicit Bernstein version of the 1D Weierstrass approximation theorem\<close>
-
-lemma Bernstein_Weierstrass:
-  fixes f :: "real \<Rightarrow> real"
-  assumes contf: "continuous_on {0..1} f" and e: "0 < e"
-    shows "\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> {0..1}
-                    \<longrightarrow> \<bar>f x - (\<Sum>k = 0..n. f(k/n) * Bernstein n k x)\<bar> < e"
-proof -
-  have "bounded (f ` {0..1})"
-    using compact_continuous_image compact_imp_bounded contf by blast
-  then obtain M where M: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<bar>f x\<bar> \<le> M"
-    by (force simp add: bounded_iff)
-  then have Mge0: "0 \<le> M" by force
-  have ucontf: "uniformly_continuous_on {0..1} f"
-    using compact_uniformly_continuous contf by blast
-  then obtain d where d: "d>0" "\<And>x x'. \<lbrakk> x \<in> {0..1}; x' \<in> {0..1}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> \<bar>f x' - f x\<bar> < e/2"
-     apply (rule uniformly_continuous_onE [where e = "e/2"])
-     using e by (auto simp: dist_norm)
-  { fix n::nat and x::real
-    assume n: "Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>) \<le> n" and x: "0 \<le> x" "x \<le> 1"
-    have "0 < n" using n by simp
-    have ed0: "- (e * d\<^sup>2) < 0"
-      using e \<open>0<d\<close> by simp
-    also have "... \<le> M * 4"
-      using \<open>0\<le>M\<close> by simp
-    finally have [simp]: "real_of_int (nat \<lceil>4 * M / (e * d\<^sup>2)\<rceil>) = real_of_int \<lceil>4 * M / (e * d\<^sup>2)\<rceil>"
-      using \<open>0\<le>M\<close> e \<open>0<d\<close>
-      by (simp add: of_nat_Suc field_simps)
-    have "4*M/(e*d\<^sup>2) + 1 \<le> real (Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>))"
-      by (simp add: of_nat_Suc real_nat_ceiling_ge)
-    also have "... \<le> real n"
-      using n by (simp add: of_nat_Suc field_simps)
-    finally have nbig: "4*M/(e*d\<^sup>2) + 1 \<le> real n" .
-    have sum_bern: "(\<Sum>k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n"
-    proof -
-      have *: "\<And>a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x"
-        by (simp add: algebra_simps power2_eq_square)
-      have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
-        apply (simp add: * setsum.distrib)
-        apply (simp add: setsum_right_distrib [symmetric] mult.assoc)
-        apply (simp add: algebra_simps power2_eq_square)
-        done
-      then have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
-        by (simp add: power2_eq_square)
-      then show ?thesis
-        using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute)
-    qed
-    { fix k
-      assume k: "k \<le> n"
-      then have kn: "0 \<le> k / n" "k / n \<le> 1"
-        by (auto simp: divide_simps)
-      consider (lessd) "\<bar>x - k / n\<bar> < d" | (ged) "d \<le> \<bar>x - k / n\<bar>"
-        by linarith
-      then have "\<bar>(f x - f (k/n))\<bar> \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
-      proof cases
-        case lessd
-        then have "\<bar>(f x - f (k/n))\<bar> < e/2"
-          using d x kn by (simp add: abs_minus_commute)
-        also have "... \<le> (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)"
-          using Mge0 d by simp
-        finally show ?thesis by simp
-      next
-        case ged
-        then have dle: "d\<^sup>2 \<le> (x - k/n)\<^sup>2"
-          by (metis d(1) less_eq_real_def power2_abs power_mono)
-        have "\<bar>(f x - f (k/n))\<bar> \<le> \<bar>f x\<bar> + \<bar>f (k/n)\<bar>"
-          by (rule abs_triangle_ineq4)
-        also have "... \<le> M+M"
-          by (meson M add_mono_thms_linordered_semiring(1) kn x)
-        also have "... \<le> 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)"
-          apply simp
-          apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified])
-          using dle \<open>d>0\<close> \<open>M\<ge>0\<close> by auto
-        also have "... \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
-          using e  by simp
-        finally show ?thesis .
-        qed
-    } note * = this
-    have "\<bar>f x - (\<Sum> k = 0..n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum> k = 0..n. (f x - f(k / n)) * Bernstein n k x\<bar>"
-      by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps)
-    also have "... \<le> (\<Sum> k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
-      apply (rule order_trans [OF setsum_abs setsum_mono])
-      using *
-      apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
-      done
-    also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
-      apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern)
-      using \<open>d>0\<close> x
-      apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
-      done
-    also have "... < e"
-      apply (simp add: field_simps)
-      using \<open>d>0\<close> nbig e \<open>n>0\<close>
-      apply (simp add: divide_simps algebra_simps)
-      using ed0 by linarith
-    finally have "\<bar>f x - (\<Sum>k = 0..n. f (real k / real n) * Bernstein n k x)\<bar> < e" .
-  }
-  then show ?thesis
-    by auto
-qed
-
-
-subsection \<open>General Stone-Weierstrass theorem\<close>
-
-text\<open>Source:
-Bruno Brosowski and Frank Deutsch.
-An Elementary Proof of the Stone-Weierstrass Theorem.
-Proceedings of the American Mathematical Society
-Volume 81, Number 1, January 1981.
-DOI: 10.2307/2043993  http://www.jstor.org/stable/2043993\<close>
-
-locale function_ring_on =
-  fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
-  assumes compact: "compact s"
-  assumes continuous: "f \<in> R \<Longrightarrow> continuous_on s f"
-  assumes add: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x + g x) \<in> R"
-  assumes mult: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x * g x) \<in> R"
-  assumes const: "(\<lambda>_. c) \<in> R"
-  assumes separable: "x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<exists>f\<in>R. f x \<noteq> f y"
-
-begin
-  lemma minus: "f \<in> R \<Longrightarrow> (\<lambda>x. - f x) \<in> R"
-    by (frule mult [OF const [of "-1"]]) simp
-
-  lemma diff: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x - g x) \<in> R"
-    unfolding diff_conv_add_uminus by (metis add minus)
-
-  lemma power: "f \<in> R \<Longrightarrow> (\<lambda>x. f x ^ n) \<in> R"
-    by (induct n) (auto simp: const mult)
-
-  lemma setsum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
-    by (induct I rule: finite_induct; simp add: const add)
-
-  lemma setprod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
-    by (induct I rule: finite_induct; simp add: const mult)
-
-  definition normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
-    where "normf f \<equiv> SUP x:s. \<bar>f x\<bar>"
-
-  lemma normf_upper: "\<lbrakk>continuous_on s f; x \<in> s\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> \<le> normf f"
-    apply (simp add: normf_def)
-    apply (rule cSUP_upper, assumption)
-    by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs)
-
-  lemma normf_least: "s \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> s \<Longrightarrow> \<bar>f x\<bar> \<le> M) \<Longrightarrow> normf f \<le> M"
-    by (simp add: normf_def cSUP_least)
-
-end
-
-lemma (in function_ring_on) one:
-  assumes U: "open U" and t0: "t0 \<in> s" "t0 \<in> U" and t1: "t1 \<in> s-U"
-    shows "\<exists>V. open V \<and> t0 \<in> V \<and> s \<inter> V \<subseteq> U \<and>
-               (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. f t > 1 - e))"
-proof -
-  have "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}" if t: "t \<in> s - U" for t
-  proof -
-    have "t \<noteq> t0" using t t0 by auto
-    then obtain g where g: "g \<in> R" "g t \<noteq> g t0"
-      using separable t0  by (metis Diff_subset subset_eq t)
-    define h where [abs_def]: "h x = g x - g t0" for x
-    have "h \<in> R"
-      unfolding h_def by (fast intro: g const diff)
-    then have hsq: "(\<lambda>w. (h w)\<^sup>2) \<in> R"
-      by (simp add: power2_eq_square mult)
-    have "h t \<noteq> h t0"
-      by (simp add: h_def g)
-    then have "h t \<noteq> 0"
-      by (simp add: h_def)
-    then have ht2: "0 < (h t)^2"
-      by simp
-    also have "... \<le> normf (\<lambda>w. (h w)\<^sup>2)"
-      using t normf_upper [where x=t] continuous [OF hsq] by force
-    finally have nfp: "0 < normf (\<lambda>w. (h w)\<^sup>2)" .
-    define p where [abs_def]: "p x = (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2" for x
-    have "p \<in> R"
-      unfolding p_def by (fast intro: hsq const mult)
-    moreover have "p t0 = 0"
-      by (simp add: p_def h_def)
-    moreover have "p t > 0"
-      using nfp ht2 by (simp add: p_def)
-    moreover have "\<And>x. x \<in> s \<Longrightarrow> p x \<in> {0..1}"
-      using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def)
-    ultimately show "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}"
-      by auto
-  qed
-  then obtain pf where pf: "\<And>t. t \<in> s-U \<Longrightarrow> pf t \<in> R \<and> pf t t0 = 0 \<and> pf t t > 0"
-                   and pf01: "\<And>t. t \<in> s-U \<Longrightarrow> pf t ` s \<subseteq> {0..1}"
-    by metis
-  have com_sU: "compact (s-U)"
-    using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed)
-  have "\<And>t. t \<in> s-U \<Longrightarrow> \<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < pf t x}"
-    apply (rule open_Collect_positive)
-    by (metis pf continuous)
-  then obtain Uf where Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t) \<and> (Uf t) \<inter> s = {x\<in>s. 0 < pf t x}"
-    by metis
-  then have open_Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t)"
-    by blast
-  have tUft: "\<And>t. t \<in> s-U \<Longrightarrow> t \<in> Uf t"
-    using pf Uf by blast
-  then have *: "s-U \<subseteq> (\<Union>x \<in> s-U. Uf x)"
-    by blast
-  obtain subU where subU: "subU \<subseteq> s - U" "finite subU" "s - U \<subseteq> (\<Union>x \<in> subU. Uf x)"
-    by (blast intro: that open_Uf compactE_image [OF com_sU _ *])
-  then have [simp]: "subU \<noteq> {}"
-    using t1 by auto
-  then have cardp: "card subU > 0" using subU
-    by (simp add: card_gt_0_iff)
-  define p where [abs_def]: "p x = (1 / card subU) * (\<Sum>t \<in> subU. pf t x)" for x
-  have pR: "p \<in> R"
-    unfolding p_def using subU pf by (fast intro: pf const mult setsum)
-  have pt0 [simp]: "p t0 = 0"
-    using subU pf by (auto simp: p_def intro: setsum.neutral)
-  have pt_pos: "p t > 0" if t: "t \<in> s-U" for t
-  proof -
-    obtain i where i: "i \<in> subU" "t \<in> Uf i" using subU t by blast
-    show ?thesis
-      using subU i t
-      apply (clarsimp simp: p_def divide_simps)
-      apply (rule setsum_pos2 [OF \<open>finite subU\<close>])
-      using Uf t pf01 apply auto
-      apply (force elim!: subsetCE)
-      done
-  qed
-  have p01: "p x \<in> {0..1}" if t: "x \<in> s" for x
-  proof -
-    have "0 \<le> p x"
-      using subU cardp t
-      apply (simp add: p_def divide_simps setsum_nonneg)
-      apply (rule setsum_nonneg)
-      using pf01 by force
-    moreover have "p x \<le> 1"
-      using subU cardp t
-      apply (simp add: p_def divide_simps setsum_nonneg)
-      apply (rule setsum_bounded_above [where 'a=real and K=1, simplified])
-      using pf01 by force
-    ultimately show ?thesis
-      by auto
-  qed
-  have "compact (p ` (s-U))"
-    by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR)
-  then have "open (- (p ` (s-U)))"
-    by (simp add: compact_imp_closed open_Compl)
-  moreover have "0 \<in> - (p ` (s-U))"
-    by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos)
-  ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \<subseteq> - (p ` (s-U))"
-    by (auto simp: elim!: openE)
-  then have pt_delta: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> delta0"
-    by (force simp: ball_def dist_norm dest: p01)
-  define \<delta> where "\<delta> = delta0/2"
-  have "delta0 \<le> 1" using delta0 p01 [of t1] t1
-      by (force simp: ball_def dist_norm dest: p01)
-  with delta0 have \<delta>01: "0 < \<delta>" "\<delta> < 1"
-    by (auto simp: \<delta>_def)
-  have pt_\<delta>: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> \<delta>"
-    using pt_delta delta0 by (force simp: \<delta>_def)
-  have "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. p x < \<delta>/2}"
-    by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const])
-  then obtain V where V: "open V" "V \<inter> s = {x\<in>s. p x < \<delta>/2}"
-    by blast
-  define k where "k = nat\<lfloor>1/\<delta>\<rfloor> + 1"
-  have "k>0"  by (simp add: k_def)
-  have "k-1 \<le> 1/\<delta>"
-    using \<delta>01 by (simp add: k_def)
-  with \<delta>01 have "k \<le> (1+\<delta>)/\<delta>"
-    by (auto simp: algebra_simps add_divide_distrib)
-  also have "... < 2/\<delta>"
-    using \<delta>01 by (auto simp: divide_simps)
-  finally have k2\<delta>: "k < 2/\<delta>" .
-  have "1/\<delta> < k"
-    using \<delta>01 unfolding k_def by linarith
-  with \<delta>01 k2\<delta> have k\<delta>: "1 < k*\<delta>" "k*\<delta> < 2"
-    by (auto simp: divide_simps)
-  define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t
-  have qR: "q n \<in> R" for n
-    by (simp add: q_def const diff power pR)
-  have q01: "\<And>n t. t \<in> s \<Longrightarrow> q n t \<in> {0..1}"
-    using p01 by (simp add: q_def power_le_one algebra_simps)
-  have qt0 [simp]: "\<And>n. n>0 \<Longrightarrow> q n t0 = 1"
-    using t0 pf by (simp add: q_def power_0_left)
-  { fix t and n::nat
-    assume t: "t \<in> s \<inter> V"
-    with \<open>k>0\<close> V have "k * p t < k * \<delta> / 2"
-       by force
-    then have "1 - (k * \<delta> / 2)^n \<le> 1 - (k * p t)^n"
-      using  \<open>k>0\<close> p01 t by (simp add: power_mono)
-    also have "... \<le> q n t"
-      using Bernoulli_inequality [of "- ((p t)^n)" "k^n"]
-      apply (simp add: q_def)
-      by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t)
-    finally have "1 - (k * \<delta> / 2) ^ n \<le> q n t" .
-  } note limitV = this
-  { fix t and n::nat
-    assume t: "t \<in> s - U"
-    with \<open>k>0\<close> U have "k * \<delta> \<le> k * p t"
-      by (simp add: pt_\<delta>)
-    with k\<delta> have kpt: "1 < k * p t"
-      by (blast intro: less_le_trans)
-    have ptn_pos: "0 < p t ^ n"
-      using pt_pos [OF t] by simp
-    have ptn_le: "p t ^ n \<le> 1"
-      by (meson DiffE atLeastAtMost_iff p01 power_le_one t)
-    have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n"
-      using pt_pos [OF t] \<open>k>0\<close> by (simp add: q_def)
-    also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)"
-      using pt_pos [OF t] \<open>k>0\<close>
-      apply simp
-      apply (simp only: times_divide_eq_right [symmetric])
-      apply (rule mult_left_mono [of "1::real", simplified])
-      apply (simp_all add: power_mult_distrib)
-      apply (rule zero_le_power)
-      using ptn_le by linarith
-    also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)"
-      apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]])
-      using \<open>k>0\<close> ptn_pos ptn_le
-      apply (auto simp: power_mult_distrib)
-      done
-    also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)"
-      using pt_pos [OF t] \<open>k>0\<close>
-      by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric])
-    also have "... \<le> (1/(k * (p t))^n) * 1"
-      apply (rule mult_left_mono [OF power_le_one])
-      using pt_pos \<open>k>0\<close> p01 power_le_one t apply auto
-      done
-    also have "... \<le> (1 / (k*\<delta>))^n"
-      using \<open>k>0\<close> \<delta>01  power_mono pt_\<delta> t
-      by (fastforce simp: field_simps)
-    finally have "q n t \<le> (1 / (real k * \<delta>)) ^ n " .
-  } note limitNonU = this
-  define NN
-    where "NN e = 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>" for e
-  have NN: "of_nat (NN e) > ln e / ln (real k * \<delta> / 2)"  "of_nat (NN e) > - ln e / ln (real k * \<delta>)"
-              if "0<e" for e
-      unfolding NN_def  by linarith+
-  have NN1: "\<And>e. e>0 \<Longrightarrow> (k * \<delta> / 2)^NN e < e"
-    apply (subst Transcendental.ln_less_cancel_iff [symmetric])
-      prefer 3 apply (subst ln_realpow)
-    using \<open>k>0\<close> \<open>\<delta>>0\<close> NN  k\<delta>
-    apply (force simp add: field_simps)+
-    done
-  have NN0: "\<And>e. e>0 \<Longrightarrow> (1/(k*\<delta>))^NN e < e"
-    apply (subst Transcendental.ln_less_cancel_iff [symmetric])
-      prefer 3 apply (subst ln_realpow)
-    using \<open>k>0\<close> \<open>\<delta>>0\<close> NN k\<delta>
-    apply (force simp add: field_simps ln_div)+
-    done
-  { fix t and e::real
-    assume "e>0"
-    have "t \<in> s \<inter> V \<Longrightarrow> 1 - q (NN e) t < e" "t \<in> s - U \<Longrightarrow> q (NN e) t < e"
-    proof -
-      assume t: "t \<in> s \<inter> V"
-      show "1 - q (NN e) t < e"
-        by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \<open>e>0\<close>]])
-    next
-      assume t: "t \<in> s - U"
-      show "q (NN e) t < e"
-      using  limitNonU [OF t] less_le_trans [OF NN0 [OF \<open>e>0\<close>]] not_le by blast
-    qed
-  } then have "\<And>e. e > 0 \<Longrightarrow> \<exists>f\<in>R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. 1 - e < f t)"
-    using q01
-    by (rule_tac x="\<lambda>x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR)
-  moreover have t0V: "t0 \<in> V"  "s \<inter> V \<subseteq> U"
-    using pt_\<delta> t0 U V \<delta>01  by fastforce+
-  ultimately show ?thesis using V t0V
-    by blast
-qed
-
-text\<open>Non-trivial case, with @{term A} and @{term B} both non-empty\<close>
-lemma (in function_ring_on) two_special:
-  assumes A: "closed A" "A \<subseteq> s" "a \<in> A"
-      and B: "closed B" "B \<subseteq> s" "b \<in> B"
-      and disj: "A \<inter> B = {}"
-      and e: "0 < e" "e < 1"
-    shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
-proof -
-  { fix w
-    assume "w \<in> A"
-    then have "open ( - B)" "b \<in> s" "w \<notin> B" "w \<in> s"
-      using assms by auto
-    then have "\<exists>V. open V \<and> w \<in> V \<and> s \<inter> V \<subseteq> -B \<and>
-               (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> V. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
-      using one [of "-B" w b] assms \<open>w \<in> A\<close> by simp
-  }
-  then obtain Vf where Vf:
-         "\<And>w. w \<in> A \<Longrightarrow> open (Vf w) \<and> w \<in> Vf w \<and> s \<inter> Vf w \<subseteq> -B \<and>
-                         (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
-    by metis
-  then have open_Vf: "\<And>w. w \<in> A \<Longrightarrow> open (Vf w)"
-    by blast
-  have tVft: "\<And>w. w \<in> A \<Longrightarrow> w \<in> Vf w"
-    using Vf by blast
-  then have setsum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
-    by blast
-  have com_A: "compact A" using A
-    by (metis compact compact_Int_closed inf.absorb_iff2)
-  obtain subA where subA: "subA \<subseteq> A" "finite subA" "A \<subseteq> (\<Union>x \<in> subA. Vf x)"
-    by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0])
-  then have [simp]: "subA \<noteq> {}"
-    using \<open>a \<in> A\<close> by auto
-  then have cardp: "card subA > 0" using subA
-    by (simp add: card_gt_0_iff)
-  have "\<And>w. w \<in> A \<Longrightarrow> \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e / card subA)"
-    using Vf e cardp by simp
-  then obtain ff where ff:
-         "\<And>w. w \<in> A \<Longrightarrow> ff w \<in> R \<and> ff w ` s \<subseteq> {0..1} \<and>
-                         (\<forall>x \<in> s \<inter> Vf w. ff w x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. ff w x > 1 - e / card subA)"
-    by metis
-  define pff where [abs_def]: "pff x = (\<Prod>w \<in> subA. ff w x)" for x
-  have pffR: "pff \<in> R"
-    unfolding pff_def using subA ff by (auto simp: intro: setprod)
-  moreover
-  have pff01: "pff x \<in> {0..1}" if t: "x \<in> s" for x
-  proof -
-    have "0 \<le> pff x"
-      using subA cardp t
-      apply (simp add: pff_def divide_simps setsum_nonneg)
-      apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg)
-      using ff by fastforce
-    moreover have "pff x \<le> 1"
-      using subA cardp t
-      apply (simp add: pff_def divide_simps setsum_nonneg)
-      apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
-      using ff by fastforce
-    ultimately show ?thesis
-      by auto
-  qed
-  moreover
-  { fix v x
-    assume v: "v \<in> subA" and x: "x \<in> Vf v" "x \<in> s"
-    from subA v have "pff x = ff v x * (\<Prod>w \<in> subA - {v}. ff w x)"
-      unfolding pff_def  by (metis setprod.remove)
-    also have "... \<le> ff v x * 1"
-      apply (rule Rings.ordered_semiring_class.mult_left_mono)
-      apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
-      using ff [THEN conjunct2, THEN conjunct1] v subA x
-      apply auto
-      apply (meson atLeastAtMost_iff contra_subsetD imageI)
-      apply (meson atLeastAtMost_iff contra_subsetD image_eqI)
-      using atLeastAtMost_iff by blast
-    also have "... < e / card subA"
-      using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x
-      by auto
-    also have "... \<le> e"
-      using cardp e by (simp add: divide_simps)
-    finally have "pff x < e" .
-  }
-  then have "\<And>x. x \<in> A \<Longrightarrow> pff x < e"
-    using A Vf subA by (metis UN_E contra_subsetD)
-  moreover
-  { fix x
-    assume x: "x \<in> B"
-    then have "x \<in> s"
-      using B by auto
-    have "1 - e \<le> (1 - e / card subA) ^ card subA"
-      using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp
-      by (auto simp: field_simps)
-    also have "... = (\<Prod>w \<in> subA. 1 - e / card subA)"
-      by (simp add: setprod_constant subA(2))
-    also have "... < pff x"
-      apply (simp add: pff_def)
-      apply (rule setprod_mono_strict [where f = "\<lambda>x. 1 - e / card subA", simplified])
-      apply (simp_all add: subA(2))
-      apply (intro ballI conjI)
-      using e apply (force simp: divide_simps)
-      using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x
-      apply blast
-      done
-    finally have "1 - e < pff x" .
-  }
-  ultimately
-  show ?thesis by blast
-qed
-
-lemma (in function_ring_on) two:
-  assumes A: "closed A" "A \<subseteq> s"
-      and B: "closed B" "B \<subseteq> s"
-      and disj: "A \<inter> B = {}"
-      and e: "0 < e" "e < 1"
-    shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
-proof (cases "A \<noteq> {} \<and> B \<noteq> {}")
-  case True then show ?thesis
-    apply (simp add: ex_in_conv [symmetric])
-    using assms
-    apply safe
-    apply (force simp add: intro!: two_special)
-    done
-next
-  case False with e show ?thesis
-    apply simp
-    apply (erule disjE)
-    apply (rule_tac [2] x="\<lambda>x. 0" in bexI)
-    apply (rule_tac x="\<lambda>x. 1" in bexI)
-    apply (auto simp: const)
-    done
-qed
-
-text\<open>The special case where @{term f} is non-negative and @{term"e<1/3"}\<close>
-lemma (in function_ring_on) Stone_Weierstrass_special:
-  assumes f: "continuous_on s f" and fpos: "\<And>x. x \<in> s \<Longrightarrow> f x \<ge> 0"
-      and e: "0 < e" "e < 1/3"
-  shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < 2*e"
-proof -
-  define n where "n = 1 + nat \<lceil>normf f / e\<rceil>"
-  define A where "A j = {x \<in> s. f x \<le> (j - 1/3)*e}" for j :: nat
-  define B where "B j = {x \<in> s. f x \<ge> (j + 1/3)*e}" for j :: nat
-  have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
-    using e
-    apply (simp_all add: n_def field_simps of_nat_Suc)
-    by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq)
-  then have ge_fx: "(n-1) * e \<ge> f x" if "x \<in> s" for x
-    using f normf_upper that by fastforce
-  { fix j
-    have A: "closed (A j)" "A j \<subseteq> s"
-      apply (simp_all add: A_def Collect_restrict)
-      apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const])
-      apply (simp add: compact compact_imp_closed)
-      done
-    have B: "closed (B j)" "B j \<subseteq> s"
-      apply (simp_all add: B_def Collect_restrict)
-      apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f])
-      apply (simp add: compact compact_imp_closed)
-      done
-    have disj: "(A j) \<inter> (B j) = {}"
-      using e by (auto simp: A_def B_def field_simps)
-    have "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A j. f x < e/n) \<and> (\<forall>x \<in> B j. f x > 1 - e/n)"
-      apply (rule two)
-      using e A B disj ngt
-      apply simp_all
-      done
-  }
-  then obtain xf where xfR: "\<And>j. xf j \<in> R" and xf01: "\<And>j. xf j ` s \<subseteq> {0..1}"
-                   and xfA: "\<And>x j. x \<in> A j \<Longrightarrow> xf j x < e/n"
-                   and xfB: "\<And>x j. x \<in> B j \<Longrightarrow> xf j x > 1 - e/n"
-    by metis
-  define g where [abs_def]: "g x = e * (\<Sum>i\<le>n. xf i x)" for x
-  have gR: "g \<in> R"
-    unfolding g_def by (fast intro: mult const setsum xfR)
-  have gge0: "\<And>x. x \<in> s \<Longrightarrow> g x \<ge> 0"
-    using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
-  have A0: "A 0 = {}"
-    using fpos e by (fastforce simp: A_def)
-  have An: "A n = s"
-    using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff)
-  have Asub: "A j \<subseteq> A i" if "i\<ge>j" for i j
-    using e that apply (clarsimp simp: A_def)
-    apply (erule order_trans, simp)
-    done
-  { fix t
-    assume t: "t \<in> s"
-    define j where "j = (LEAST j. t \<in> A j)"
-    have jn: "j \<le> n"
-      using t An by (simp add: Least_le j_def)
-    have Aj: "t \<in> A j"
-      using t An by (fastforce simp add: j_def intro: LeastI)
-    then have Ai: "t \<in> A i" if "i\<ge>j" for i
-      using Asub [OF that] by blast
-    then have fj1: "f t \<le> (j - 1/3)*e"
-      by (simp add: A_def)
-    then have Anj: "t \<notin> A i" if "i<j" for i
-      using  Aj  \<open>i<j\<close>
-      apply (simp add: j_def)
-      using not_less_Least by blast
-    have j1: "1 \<le> j"
-      using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def)
-    then have Anj: "t \<notin> A (j-1)"
-      using Least_le by (fastforce simp add: j_def)
-    then have fj2: "(j - 4/3)*e < f t"
-      using j1 t  by (simp add: A_def of_nat_diff)
-    have ***: "xf i t \<le> e/n" if "i\<ge>j" for i
-      using xfA [OF Ai] that by (simp add: less_eq_real_def)
-    { fix i
-      assume "i+2 \<le> j"
-      then obtain d where "i+2+d = j"
-        using le_Suc_ex that by blast
-      then have "t \<in> B i"
-        using Anj e ge_fx [OF t] \<open>1 \<le> n\<close> fpos [OF t] t
-        apply (simp add: A_def B_def)
-        apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc)
-        apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"])
-        apply auto
-        done
-      then have "xf i t > 1 - e/n"
-        by (rule xfB)
-    } note **** = this
-    have xf_le1: "\<And>i. xf i t \<le> 1"
-      using xf01 t by force
-    have "g t = e * (\<Sum>i<j. xf i t) + e * (\<Sum>i=j..n. xf i t)"
-      using j1 jn e
-      apply (simp add: g_def distrib_left [symmetric])
-      apply (subst setsum.union_disjoint [symmetric])
-      apply (auto simp: ivl_disj_un)
-      done
-    also have "... \<le> e*j + e * ((Suc n - j)*e/n)"
-      apply (rule add_mono)
-      apply (simp_all only: mult_le_cancel_left_pos e)
-      apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
-      using setsum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
-      apply simp
-      done
-    also have "... \<le> j*e + e*(n - j + 1)*e/n "
-      using \<open>1 \<le> n\<close> e  by (simp add: field_simps del: of_nat_Suc)
-    also have "... \<le> j*e + e*e"
-      using \<open>1 \<le> n\<close> e j1 by (simp add: field_simps del: of_nat_Suc)
-    also have "... < (j + 1/3)*e"
-      using e by (auto simp: field_simps)
-    finally have gj1: "g t < (j + 1 / 3) * e" .
-    have gj2: "(j - 4/3)*e < g t"
-    proof (cases "2 \<le> j")
-      case False
-      then have "j=1" using j1 by simp
-      with t gge0 e show ?thesis by force
-    next
-      case True
-      then have "(j - 4/3)*e < (j-1)*e - e^2"
-        using e by (auto simp: of_nat_diff algebra_simps power2_eq_square)
-      also have "... < (j-1)*e - ((j - 1)/n) * e^2"
-        using e True jn by (simp add: power2_eq_square field_simps)
-      also have "... = e * (j-1) * (1 - e/n)"
-        by (simp add: power2_eq_square field_simps)
-      also have "... \<le> e * (\<Sum>i\<le>j-2. xf i t)"
-        using e
-        apply simp
-        apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]])
-        using True
-        apply (simp_all add: of_nat_Suc of_nat_diff)
-        done
-      also have "... \<le> g t"
-        using jn e
-        using e xf01 t
-        apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
-        apply (rule Groups_Big.setsum_mono2, auto)
-        done
-      finally show ?thesis .
-    qed
-    have "\<bar>f t - g t\<bar> < 2 * e"
-      using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps)
-  }
-  then show ?thesis
-    by (rule_tac x=g in bexI) (auto intro: gR)
-qed
-
-text\<open>The ``unpretentious'' formulation\<close>
-lemma (in function_ring_on) Stone_Weierstrass_basic:
-  assumes f: "continuous_on s f" and e: "e > 0"
-  shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < e"
-proof -
-  have "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>(f x + normf f) - g x\<bar> < 2 * min (e/2) (1/4)"
-    apply (rule Stone_Weierstrass_special)
-    apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const])
-    using normf_upper [OF f] apply force
-    apply (simp add: e, linarith)
-    done
-  then obtain g where "g \<in> R" "\<forall>x\<in>s. \<bar>g x - (f x + normf f)\<bar> < e"
-    by force
-  then show ?thesis
-    apply (rule_tac x="\<lambda>x. g x - normf f" in bexI)
-    apply (auto simp: algebra_simps intro: diff const)
-    done
-qed
-
-
-theorem (in function_ring_on) Stone_Weierstrass:
-  assumes f: "continuous_on s f"
-  shows "\<exists>F\<in>UNIV \<rightarrow> R. LIM n sequentially. F n :> uniformly_on s f"
-proof -
-  { fix e::real
-    assume e: "0 < e"
-    then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e"
-      by (auto simp: real_arch_inverse [of e])
-    { fix n :: nat and x :: 'a and g :: "'a \<Rightarrow> real"
-      assume n: "N \<le> n"  "\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n)"
-      assume x: "x \<in> s"
-      have "\<not> real (Suc n) < inverse e"
-        using \<open>N \<le> n\<close> N using less_imp_inverse_less by force
-      then have "1 / (1 + real n) \<le> e"
-        using e by (simp add: field_simps of_nat_Suc)
-      then have "\<bar>f x - g x\<bar> < e"
-        using n(2) x by auto
-    } note * = this
-    have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>s. \<bar>f x - (SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n))) x\<bar> < e"
-      apply (rule eventually_sequentiallyI [of N])
-      apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *)
-      done
-  } then
-  show ?thesis
-    apply (rule_tac x="\<lambda>n::nat. SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + n))" in bexI)
-    prefer 2  apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]])
-    unfolding uniform_limit_iff
-    apply (auto simp: dist_norm abs_minus_commute)
-    done
-qed
-
-text\<open>A HOL Light formulation\<close>
-corollary Stone_Weierstrass_HOL:
-  fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
-  assumes "compact s"  "\<And>c. P(\<lambda>x. c::real)"
-          "\<And>f. P f \<Longrightarrow> continuous_on s f"
-          "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x + g x)"  "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x * g x)"
-          "\<And>x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) \<Longrightarrow> \<exists>f. P(f) \<and> ~(f x = f y)"
-          "continuous_on s f"
-       "0 < e"
-    shows "\<exists>g. P(g) \<and> (\<forall>x \<in> s. \<bar>f x - g x\<bar> < e)"
-proof -
-  interpret PR: function_ring_on "Collect P"
-    apply unfold_locales
-    using assms
-    by auto
-  show ?thesis
-    using PR.Stone_Weierstrass_basic [OF \<open>continuous_on s f\<close> \<open>0 < e\<close>]
-    by blast
-qed
-
-
-subsection \<open>Polynomial functions\<close>
-
-inductive real_polynomial_function :: "('a::real_normed_vector \<Rightarrow> real) \<Rightarrow> bool" where
-    linear: "bounded_linear f \<Longrightarrow> real_polynomial_function f"
-  | const: "real_polynomial_function (\<lambda>x. c)"
-  | add:   "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x + g x)"
-  | mult:  "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x * g x)"
-
-declare real_polynomial_function.intros [intro]
-
-definition polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
-  where
-   "polynomial_function p \<equiv> (\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f o p))"
-
-lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p"
-unfolding polynomial_function_def
-proof
-  assume "real_polynomial_function p"
-  then show " \<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
-  proof (induction p rule: real_polynomial_function.induct)
-    case (linear h) then show ?case
-      by (auto simp: bounded_linear_compose real_polynomial_function.linear)
-  next
-    case (const h) then show ?case
-      by (simp add: real_polynomial_function.const)
-  next
-    case (add h) then show ?case
-      by (force simp add: bounded_linear_def linear_add real_polynomial_function.add)
-  next
-    case (mult h) then show ?case
-      by (force simp add: real_bounded_linear const real_polynomial_function.mult)
-  qed
-next
-  assume [rule_format, OF bounded_linear_ident]: "\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
-  then show "real_polynomial_function p"
-    by (simp add: o_def)
-qed
-
-lemma polynomial_function_const [iff]: "polynomial_function (\<lambda>x. c)"
-  by (simp add: polynomial_function_def o_def const)
-
-lemma polynomial_function_bounded_linear:
-  "bounded_linear f \<Longrightarrow> polynomial_function f"
-  by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear)
-
-lemma polynomial_function_id [iff]: "polynomial_function(\<lambda>x. x)"
-  by (simp add: polynomial_function_bounded_linear)
-
-lemma polynomial_function_add [intro]:
-    "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x + g x)"
-  by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def)
-
-lemma polynomial_function_mult [intro]:
-  assumes f: "polynomial_function f" and g: "polynomial_function g"
-    shows "polynomial_function (\<lambda>x. f x *\<^sub>R g x)"
-  using g
-  apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR  const real_polynomial_function.mult o_def)
-  apply (rule mult)
-  using f
-  apply (auto simp: real_polynomial_function_eq)
-  done
-
-lemma polynomial_function_cmul [intro]:
-  assumes f: "polynomial_function f"
-    shows "polynomial_function (\<lambda>x. c *\<^sub>R f x)"
-  by (rule polynomial_function_mult [OF polynomial_function_const f])
-
-lemma polynomial_function_minus [intro]:
-  assumes f: "polynomial_function f"
-    shows "polynomial_function (\<lambda>x. - f x)"
-  using polynomial_function_cmul [OF f, of "-1"] by simp
-
-lemma polynomial_function_diff [intro]:
-    "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x - g x)"
-  unfolding add_uminus_conv_diff [symmetric]
-  by (metis polynomial_function_add polynomial_function_minus)
-
-lemma polynomial_function_setsum [intro]:
-    "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. setsum (f x) I)"
-by (induct I rule: finite_induct) auto
-
-lemma real_polynomial_function_minus [intro]:
-    "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. - f x)"
-  using polynomial_function_minus [of f]
-  by (simp add: real_polynomial_function_eq)
-
-lemma real_polynomial_function_diff [intro]:
-    "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x - g x)"
-  using polynomial_function_diff [of f]
-  by (simp add: real_polynomial_function_eq)
-
-lemma real_polynomial_function_setsum [intro]:
-    "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. setsum (f x) I)"
-  using polynomial_function_setsum [of I f]
-  by (simp add: real_polynomial_function_eq)
-
-lemma real_polynomial_function_power [intro]:
-    "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. f x ^ n)"
-  by (induct n) (simp_all add: const mult)
-
-lemma real_polynomial_function_compose [intro]:
-  assumes f: "polynomial_function f" and g: "real_polynomial_function g"
-    shows "real_polynomial_function (g o f)"
-  using g
-  apply (induction g rule: real_polynomial_function.induct)
-  using f
-  apply (simp_all add: polynomial_function_def o_def const add mult)
-  done
-
-lemma polynomial_function_compose [intro]:
-  assumes f: "polynomial_function f" and g: "polynomial_function g"
-    shows "polynomial_function (g o f)"
-  using g real_polynomial_function_compose [OF f]
-  by (auto simp: polynomial_function_def o_def)
-
-lemma setsum_max_0:
-  fixes x::real (*in fact "'a::comm_ring_1"*)
-  shows "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..m. x^i * a i)"
-proof -
-  have "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..max m n. (if i \<le> m then x^i * a i else 0))"
-    by (auto simp: algebra_simps intro: setsum.cong)
-  also have "... = (\<Sum>i = 0..m. (if i \<le> m then x^i * a i else 0))"
-    by (rule setsum.mono_neutral_right) auto
-  also have "... = (\<Sum>i = 0..m. x^i * a i)"
-    by (auto simp: algebra_simps intro: setsum.cong)
-  finally show ?thesis .
-qed
-
-lemma real_polynomial_function_imp_setsum:
-  assumes "real_polynomial_function f"
-    shows "\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i)"
-using assms
-proof (induct f)
-  case (linear f)
-  then show ?case
-    apply (clarsimp simp add: real_bounded_linear)
-    apply (rule_tac x="\<lambda>i. if i=0 then 0 else c" in exI)
-    apply (rule_tac x=1 in exI)
-    apply (simp add: mult_ac)
-    done
-next
-  case (const c)
-  show ?case
-    apply (rule_tac x="\<lambda>i. c" in exI)
-    apply (rule_tac x=0 in exI)
-    apply (auto simp: mult_ac of_nat_Suc)
-    done
-  case (add f1 f2)
-  then obtain a1 n1 a2 n2 where
-    "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
-    by auto
-  then show ?case
-    apply (rule_tac x="\<lambda>i. (if i \<le> n1 then a1 i else 0) + (if i \<le> n2 then a2 i else 0)" in exI)
-    apply (rule_tac x="max n1 n2" in exI)
-    using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1]
-    apply (simp add: setsum.distrib algebra_simps max.commute)
-    done
-  case (mult f1 f2)
-  then obtain a1 n1 a2 n2 where
-    "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
-    by auto
-  then obtain b1 b2 where
-    "f1 = (\<lambda>x. \<Sum>i = 0..n1. b1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. b2 i * x ^ i)"
-    "b1 = (\<lambda>i. if i\<le>n1 then a1 i else 0)" "b2 = (\<lambda>i. if i\<le>n2 then a2 i else 0)"
-    by auto
-  then show ?case
-    apply (rule_tac x="\<lambda>i. \<Sum>k\<le>i. b1 k * b2 (i - k)" in exI)
-    apply (rule_tac x="n1+n2" in exI)
-    using polynomial_product [of n1 b1 n2 b2]
-    apply (simp add: Set_Interval.atLeast0AtMost)
-    done
-qed
-
-lemma real_polynomial_function_iff_setsum:
-     "real_polynomial_function f \<longleftrightarrow> (\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i))"
-  apply (rule iffI)
-  apply (erule real_polynomial_function_imp_setsum)
-  apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum)
-  done
-
-lemma polynomial_function_iff_Basis_inner:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
-  shows "polynomial_function f \<longleftrightarrow> (\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. inner (f x) b))"
-        (is "?lhs = ?rhs")
-unfolding polynomial_function_def
-proof (intro iffI allI impI)
-  assume "\<forall>h. bounded_linear h \<longrightarrow> real_polynomial_function (h \<circ> f)"
-  then show ?rhs
-    by (force simp add: bounded_linear_inner_left o_def)
-next
-  fix h :: "'b \<Rightarrow> real"
-  assume rp: "\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. f x \<bullet> b)" and h: "bounded_linear h"
-  have "real_polynomial_function (h \<circ> (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b))"
-    apply (rule real_polynomial_function_compose [OF _  linear [OF h]])
-    using rp
-    apply (auto simp: real_polynomial_function_eq polynomial_function_mult)
-    done
-  then show "real_polynomial_function (h \<circ> f)"
-    by (simp add: euclidean_representation_setsum_fun)
-qed
-
-subsection \<open>Stone-Weierstrass theorem for polynomial functions\<close>
-
-text\<open>First, we need to show that they are continous, differentiable and separable.\<close>
-
-lemma continuous_real_polymonial_function:
-  assumes "real_polynomial_function f"
-    shows "continuous (at x) f"
-using assms
-by (induct f) (auto simp: linear_continuous_at)
-
-lemma continuous_polymonial_function:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
-  assumes "polynomial_function f"
-    shows "continuous (at x) f"
-  apply (rule euclidean_isCont)
-  using assms apply (simp add: polynomial_function_iff_Basis_inner)
-  apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR)
-  done
-
-lemma continuous_on_polymonial_function:
-  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
-  assumes "polynomial_function f"
-    shows "continuous_on s f"
-  using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on
-  by blast
-
-lemma has_real_derivative_polynomial_function:
-  assumes "real_polynomial_function p"
-    shows "\<exists>p'. real_polynomial_function p' \<and>
-                 (\<forall>x. (p has_real_derivative (p' x)) (at x))"
-using assms
-proof (induct p)
-  case (linear p)
-  then show ?case
-    by (force simp: real_bounded_linear const intro!: derivative_eq_intros)
-next
-  case (const c)
-  show ?case
-    by (rule_tac x="\<lambda>x. 0" in exI) auto
-  case (add f1 f2)
-  then obtain p1 p2 where
-    "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
-    "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
-    by auto
-  then show ?case
-    apply (rule_tac x="\<lambda>x. p1 x + p2 x" in exI)
-    apply (auto intro!: derivative_eq_intros)
-    done
-  case (mult f1 f2)
-  then obtain p1 p2 where
-    "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
-    "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
-    by auto
-  then show ?case
-    using mult
-    apply (rule_tac x="\<lambda>x. f1 x * p2 x + f2 x * p1 x" in exI)
-    apply (auto intro!: derivative_eq_intros)
-    done
-qed
-
-lemma has_vector_derivative_polynomial_function:
-  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "polynomial_function p"
-    shows "\<exists>p'. polynomial_function p' \<and>
-                 (\<forall>x. (p has_vector_derivative (p' x)) (at x))"
-proof -
-  { fix b :: 'a
-    assume "b \<in> Basis"
-    then
-    obtain p' where p': "real_polynomial_function p'" and pd: "\<And>x. ((\<lambda>x. p x \<bullet> b) has_real_derivative p' x) (at x)"
-      using assms [unfolded polynomial_function_iff_Basis_inner, rule_format]  \<open>b \<in> Basis\<close>
-      has_real_derivative_polynomial_function
-      by blast
-    have "\<exists>q. polynomial_function q \<and> (\<forall>x. ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative q x) (at x))"
-      apply (rule_tac x="\<lambda>x. p' x *\<^sub>R b" in exI)
-      using \<open>b \<in> Basis\<close> p'
-      apply (simp add: polynomial_function_iff_Basis_inner inner_Basis)
-      apply (auto intro: derivative_eq_intros pd)
-      done
-  }
-  then obtain qf where qf:
-      "\<And>b. b \<in> Basis \<Longrightarrow> polynomial_function (qf b)"
-      "\<And>b x. b \<in> Basis \<Longrightarrow> ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative qf b x) (at x)"
-    by metis
-  show ?thesis
-    apply (subst euclidean_representation_setsum_fun [of p, symmetric])
-    apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. qf b x" in exI)
-    apply (auto intro: has_vector_derivative_setsum qf)
-    done
-qed
-
-lemma real_polynomial_function_separable:
-  fixes x :: "'a::euclidean_space"
-  assumes "x \<noteq> y" shows "\<exists>f. real_polynomial_function f \<and> f x \<noteq> f y"
-proof -
-  have "real_polynomial_function (\<lambda>u. \<Sum>b\<in>Basis. (inner (x-u) b)^2)"
-    apply (rule real_polynomial_function_setsum)
-    apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff
-                 const linear bounded_linear_inner_left)
-    done
-  then show ?thesis
-    apply (intro exI conjI, assumption)
-    using assms
-    apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps)
-    done
-qed
-
-lemma Stone_Weierstrass_real_polynomial_function:
-  fixes f :: "'a::euclidean_space \<Rightarrow> real"
-  assumes "compact s" "continuous_on s f" "0 < e"
-    shows "\<exists>g. real_polynomial_function g \<and> (\<forall>x \<in> s. \<bar>f x - g x\<bar> < e)"
-proof -
-  interpret PR: function_ring_on "Collect real_polynomial_function"
-    apply unfold_locales
-    using assms continuous_on_polymonial_function real_polynomial_function_eq
-    apply (auto intro: real_polynomial_function_separable)
-    done
-  show ?thesis
-    using PR.Stone_Weierstrass_basic [OF \<open>continuous_on s f\<close> \<open>0 < e\<close>]
-    by blast
-qed
-
-lemma Stone_Weierstrass_polynomial_function:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-  assumes s: "compact s"
-      and f: "continuous_on s f"
-      and e: "0 < e"
-    shows "\<exists>g. polynomial_function g \<and> (\<forall>x \<in> s. norm(f x - g x) < e)"
-proof -
-  { fix b :: 'b
-    assume "b \<in> Basis"
-    have "\<exists>p. real_polynomial_function p \<and> (\<forall>x \<in> s. \<bar>f x \<bullet> b - p x\<bar> < e / DIM('b))"
-      apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF s _, of "\<lambda>x. f x \<bullet> b" "e / card Basis"]])
-      using e f
-      apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros)
-      done
-  }
-  then obtain pf where pf:
-      "\<And>b. b \<in> Basis \<Longrightarrow> real_polynomial_function (pf b) \<and> (\<forall>x \<in> s. \<bar>f x \<bullet> b - pf b x\<bar> < e / DIM('b))"
-      apply (rule bchoice [rule_format, THEN exE])
-      apply assumption
-      apply (force simp add: intro: that)
-      done
-  have "polynomial_function (\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b)"
-    using pf
-    by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq)
-  moreover
-  { fix x
-    assume "x \<in> s"
-    have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) \<le> (\<Sum>b\<in>Basis. norm ((f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b))"
-      by (rule norm_setsum)
-    also have "... < of_nat DIM('b) * (e / DIM('b))"
-      apply (rule setsum_bounded_above_strict)
-      apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \<open>x \<in> s\<close>)
-      apply (rule DIM_positive)
-      done
-    also have "... = e"
-      using DIM_positive by (simp add: field_simps)
-    finally have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) < e" .
-  }
-  ultimately
-  show ?thesis
-    apply (subst euclidean_representation_setsum_fun [of f, symmetric])
-    apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b" in exI)
-    apply (auto simp: setsum_subtractf [symmetric])
-    done
-qed
-
-
-subsection\<open>Polynomial functions as paths\<close>
-
-text\<open>One application is to pick a smooth approximation to a path,
-or just pick a smooth path anyway in an open connected set\<close>
-
-lemma path_polynomial_function:
-    fixes g  :: "real \<Rightarrow> 'b::euclidean_space"
-    shows "polynomial_function g \<Longrightarrow> path g"
-  by (simp add: path_def continuous_on_polymonial_function)
-
-lemma path_approx_polynomial_function:
-    fixes g :: "real \<Rightarrow> 'b::euclidean_space"
-    assumes "path g" "0 < e"
-    shows "\<exists>p. polynomial_function p \<and>
-                pathstart p = pathstart g \<and>
-                pathfinish p = pathfinish g \<and>
-                (\<forall>t \<in> {0..1}. norm(p t - g t) < e)"
-proof -
-  obtain q where poq: "polynomial_function q" and noq: "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (g x - q x) < e/4"
-    using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms
-    by (auto simp: path_def)
-  have pf: "polynomial_function (\<lambda>t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))"
-    by (force simp add: poq)
-  have *: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)"
-    apply (intro Real_Vector_Spaces.norm_add_less)
-    using noq
-    apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1)
-    done
-  show ?thesis
-    apply (intro exI conjI)
-    apply (rule pf)
-    using *
-    apply (auto simp add: pathstart_def pathfinish_def algebra_simps)
-    done
-qed
-
-lemma connected_open_polynomial_connected:
-  fixes s :: "'a::euclidean_space set"
-  assumes s: "open s" "connected s"
-      and "x \<in> s" "y \<in> s"
-    shows "\<exists>g. polynomial_function g \<and> path_image g \<subseteq> s \<and>
-               pathstart g = x \<and> pathfinish g = y"
-proof -
-  have "path_connected s" using assms
-    by (simp add: connected_open_path_connected)
-  with \<open>x \<in> s\<close> \<open>y \<in> s\<close> obtain p where p: "path p" "path_image p \<subseteq> s" "pathstart p = x" "pathfinish p = y"
-    by (force simp: path_connected_def)
-  have "\<exists>e. 0 < e \<and> (\<forall>x \<in> path_image p. ball x e \<subseteq> s)"
-  proof (cases "s = UNIV")
-    case True then show ?thesis
-      by (simp add: gt_ex)
-  next
-    case False
-    then have "- s \<noteq> {}" by blast
-    then show ?thesis
-      apply (rule_tac x="setdist (path_image p) (-s)" in exI)
-      using s p
-      apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed)
-      using setdist_le_dist [of _ "path_image p" _ "-s"]
-      by fastforce
-  qed
-  then obtain e where "0 < e"and eb: "\<And>x. x \<in> path_image p \<Longrightarrow> ball x e \<subseteq> s"
-    by auto
-  show ?thesis
-    using path_approx_polynomial_function [OF \<open>path p\<close> \<open>0 < e\<close>]
-    apply clarify
-    apply (intro exI conjI, assumption)
-    using p
-    apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+
-    done
-qed
-
-hide_fact linear add mult const
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Weierstrass_Theorems.thy	Thu Aug 04 19:36:31 2016 +0200
@@ -0,0 +1,1218 @@
+section \<open>The Bernstein-Weierstrass and Stone-Weierstrass Theorems\<close>
+
+text\<open>By L C Paulson (2015)\<close>
+
+theory Weierstrass_Theorems
+imports Uniform_Limit Path_Connected
+begin
+
+subsection \<open>Bernstein polynomials\<close>
+
+definition Bernstein :: "[nat,nat,real] \<Rightarrow> real" where
+  "Bernstein n k x \<equiv> of_nat (n choose k) * x ^ k * (1 - x) ^ (n - k)"
+
+lemma Bernstein_nonneg: "\<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> Bernstein n k x"
+  by (simp add: Bernstein_def)
+
+lemma Bernstein_pos: "\<lbrakk>0 < x; x < 1; k \<le> n\<rbrakk> \<Longrightarrow> 0 < Bernstein n k x"
+  by (simp add: Bernstein_def)
+
+lemma sum_Bernstein [simp]: "(\<Sum> k = 0..n. Bernstein n k x) = 1"
+  using binomial_ring [of x "1-x" n]
+  by (simp add: Bernstein_def)
+
+lemma binomial_deriv1:
+    "(\<Sum>k=0..n. (of_nat k * of_nat (n choose k)) * a^(k-1) * b^(n-k)) = real_of_nat n * (a+b) ^ (n-1)"
+  apply (rule DERIV_unique [where f = "\<lambda>a. (a+b)^n" and x=a])
+  apply (subst binomial_ring)
+  apply (rule derivative_eq_intros setsum.cong | simp)+
+  done
+
+lemma binomial_deriv2:
+    "(\<Sum>k=0..n. (of_nat k * of_nat (k-1) * of_nat (n choose k)) * a^(k-2) * b^(n-k)) =
+     of_nat n * of_nat (n-1) * (a+b::real) ^ (n-2)"
+  apply (rule DERIV_unique [where f = "\<lambda>a. of_nat n * (a+b::real) ^ (n-1)" and x=a])
+  apply (subst binomial_deriv1 [symmetric])
+  apply (rule derivative_eq_intros setsum.cong | simp add: Num.numeral_2_eq_2)+
+  done
+
+lemma sum_k_Bernstein [simp]: "(\<Sum>k = 0..n. real k * Bernstein n k x) = of_nat n * x"
+  apply (subst binomial_deriv1 [of n x "1-x", simplified, symmetric])
+  apply (simp add: setsum_left_distrib)
+  apply (auto simp: Bernstein_def algebra_simps realpow_num_eq_if intro!: setsum.cong)
+  done
+
+lemma sum_kk_Bernstein [simp]: "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real n * (real n - 1) * x\<^sup>2"
+proof -
+  have "(\<Sum> k = 0..n. real k * (real k - 1) * Bernstein n k x) = real_of_nat n * real_of_nat (n - Suc 0) * x\<^sup>2"
+    apply (subst binomial_deriv2 [of n x "1-x", simplified, symmetric])
+    apply (simp add: setsum_left_distrib)
+    apply (rule setsum.cong [OF refl])
+    apply (simp add: Bernstein_def power2_eq_square algebra_simps)
+    apply (rename_tac k)
+    apply (subgoal_tac "k = 0 \<or> k = 1 \<or> (\<exists>k'. k = Suc (Suc k'))")
+    apply (force simp add: field_simps of_nat_Suc power2_eq_square)
+    by presburger
+  also have "... = n * (n - 1) * x\<^sup>2"
+    by auto
+  finally show ?thesis
+    by auto
+qed
+
+subsection \<open>Explicit Bernstein version of the 1D Weierstrass approximation theorem\<close>
+
+lemma Bernstein_Weierstrass:
+  fixes f :: "real \<Rightarrow> real"
+  assumes contf: "continuous_on {0..1} f" and e: "0 < e"
+    shows "\<exists>N. \<forall>n x. N \<le> n \<and> x \<in> {0..1}
+                    \<longrightarrow> \<bar>f x - (\<Sum>k = 0..n. f(k/n) * Bernstein n k x)\<bar> < e"
+proof -
+  have "bounded (f ` {0..1})"
+    using compact_continuous_image compact_imp_bounded contf by blast
+  then obtain M where M: "\<And>x. 0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> \<bar>f x\<bar> \<le> M"
+    by (force simp add: bounded_iff)
+  then have Mge0: "0 \<le> M" by force
+  have ucontf: "uniformly_continuous_on {0..1} f"
+    using compact_uniformly_continuous contf by blast
+  then obtain d where d: "d>0" "\<And>x x'. \<lbrakk> x \<in> {0..1}; x' \<in> {0..1}; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> \<bar>f x' - f x\<bar> < e/2"
+     apply (rule uniformly_continuous_onE [where e = "e/2"])
+     using e by (auto simp: dist_norm)
+  { fix n::nat and x::real
+    assume n: "Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>) \<le> n" and x: "0 \<le> x" "x \<le> 1"
+    have "0 < n" using n by simp
+    have ed0: "- (e * d\<^sup>2) < 0"
+      using e \<open>0<d\<close> by simp
+    also have "... \<le> M * 4"
+      using \<open>0\<le>M\<close> by simp
+    finally have [simp]: "real_of_int (nat \<lceil>4 * M / (e * d\<^sup>2)\<rceil>) = real_of_int \<lceil>4 * M / (e * d\<^sup>2)\<rceil>"
+      using \<open>0\<le>M\<close> e \<open>0<d\<close>
+      by (simp add: of_nat_Suc field_simps)
+    have "4*M/(e*d\<^sup>2) + 1 \<le> real (Suc (nat\<lceil>4*M/(e*d\<^sup>2)\<rceil>))"
+      by (simp add: of_nat_Suc real_nat_ceiling_ge)
+    also have "... \<le> real n"
+      using n by (simp add: of_nat_Suc field_simps)
+    finally have nbig: "4*M/(e*d\<^sup>2) + 1 \<le> real n" .
+    have sum_bern: "(\<Sum>k = 0..n. (x - k/n)\<^sup>2 * Bernstein n k x) = x * (1 - x) / n"
+    proof -
+      have *: "\<And>a b x::real. (a - b)\<^sup>2 * x = a * (a - 1) * x + (1 - 2 * b) * a * x + b * b * x"
+        by (simp add: algebra_simps power2_eq_square)
+      have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x) = n * x * (1 - x)"
+        apply (simp add: * setsum.distrib)
+        apply (simp add: setsum_right_distrib [symmetric] mult.assoc)
+        apply (simp add: algebra_simps power2_eq_square)
+        done
+      then have "(\<Sum> k = 0..n. (k - n * x)\<^sup>2 * Bernstein n k x)/n^2 = x * (1 - x) / n"
+        by (simp add: power2_eq_square)
+      then show ?thesis
+        using n by (simp add: setsum_divide_distrib divide_simps mult.commute power2_commute)
+    qed
+    { fix k
+      assume k: "k \<le> n"
+      then have kn: "0 \<le> k / n" "k / n \<le> 1"
+        by (auto simp: divide_simps)
+      consider (lessd) "\<bar>x - k / n\<bar> < d" | (ged) "d \<le> \<bar>x - k / n\<bar>"
+        by linarith
+      then have "\<bar>(f x - f (k/n))\<bar> \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
+      proof cases
+        case lessd
+        then have "\<bar>(f x - f (k/n))\<bar> < e/2"
+          using d x kn by (simp add: abs_minus_commute)
+        also have "... \<le> (e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2)"
+          using Mge0 d by simp
+        finally show ?thesis by simp
+      next
+        case ged
+        then have dle: "d\<^sup>2 \<le> (x - k/n)\<^sup>2"
+          by (metis d(1) less_eq_real_def power2_abs power_mono)
+        have "\<bar>(f x - f (k/n))\<bar> \<le> \<bar>f x\<bar> + \<bar>f (k/n)\<bar>"
+          by (rule abs_triangle_ineq4)
+        also have "... \<le> M+M"
+          by (meson M add_mono_thms_linordered_semiring(1) kn x)
+        also have "... \<le> 2 * M * ((x - k/n)\<^sup>2 / d\<^sup>2)"
+          apply simp
+          apply (rule Rings.ordered_semiring_class.mult_left_mono [of 1 "((x - k/n)\<^sup>2 / d\<^sup>2)", simplified])
+          using dle \<open>d>0\<close> \<open>M\<ge>0\<close> by auto
+        also have "... \<le> e/2 + 2 * M / d\<^sup>2 * (x - k/n)\<^sup>2"
+          using e  by simp
+        finally show ?thesis .
+        qed
+    } note * = this
+    have "\<bar>f x - (\<Sum> k = 0..n. f(k / n) * Bernstein n k x)\<bar> \<le> \<bar>\<Sum> k = 0..n. (f x - f(k / n)) * Bernstein n k x\<bar>"
+      by (simp add: setsum_subtractf setsum_right_distrib [symmetric] algebra_simps)
+    also have "... \<le> (\<Sum> k = 0..n. (e/2 + (2 * M / d\<^sup>2) * (x - k / n)\<^sup>2) * Bernstein n k x)"
+      apply (rule order_trans [OF setsum_abs setsum_mono])
+      using *
+      apply (simp add: abs_mult Bernstein_nonneg x mult_right_mono)
+      done
+    also have "... \<le> e/2 + (2 * M) / (d\<^sup>2 * n)"
+      apply (simp only: setsum.distrib Rings.semiring_class.distrib_right setsum_right_distrib [symmetric] mult.assoc sum_bern)
+      using \<open>d>0\<close> x
+      apply (simp add: divide_simps Mge0 mult_le_one mult_left_le)
+      done
+    also have "... < e"
+      apply (simp add: field_simps)
+      using \<open>d>0\<close> nbig e \<open>n>0\<close>
+      apply (simp add: divide_simps algebra_simps)
+      using ed0 by linarith
+    finally have "\<bar>f x - (\<Sum>k = 0..n. f (real k / real n) * Bernstein n k x)\<bar> < e" .
+  }
+  then show ?thesis
+    by auto
+qed
+
+
+subsection \<open>General Stone-Weierstrass theorem\<close>
+
+text\<open>Source:
+Bruno Brosowski and Frank Deutsch.
+An Elementary Proof of the Stone-Weierstrass Theorem.
+Proceedings of the American Mathematical Society
+Volume 81, Number 1, January 1981.
+DOI: 10.2307/2043993  http://www.jstor.org/stable/2043993\<close>
+
+locale function_ring_on =
+  fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
+  assumes compact: "compact s"
+  assumes continuous: "f \<in> R \<Longrightarrow> continuous_on s f"
+  assumes add: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x + g x) \<in> R"
+  assumes mult: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x * g x) \<in> R"
+  assumes const: "(\<lambda>_. c) \<in> R"
+  assumes separable: "x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<exists>f\<in>R. f x \<noteq> f y"
+
+begin
+  lemma minus: "f \<in> R \<Longrightarrow> (\<lambda>x. - f x) \<in> R"
+    by (frule mult [OF const [of "-1"]]) simp
+
+  lemma diff: "f \<in> R \<Longrightarrow> g \<in> R \<Longrightarrow> (\<lambda>x. f x - g x) \<in> R"
+    unfolding diff_conv_add_uminus by (metis add minus)
+
+  lemma power: "f \<in> R \<Longrightarrow> (\<lambda>x. f x ^ n) \<in> R"
+    by (induct n) (auto simp: const mult)
+
+  lemma setsum: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>i \<in> I. f i x) \<in> R"
+    by (induct I rule: finite_induct; simp add: const add)
+
+  lemma setprod: "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> f i \<in> R\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Prod>i \<in> I. f i x) \<in> R"
+    by (induct I rule: finite_induct; simp add: const mult)
+
+  definition normf :: "('a::t2_space \<Rightarrow> real) \<Rightarrow> real"
+    where "normf f \<equiv> SUP x:s. \<bar>f x\<bar>"
+
+  lemma normf_upper: "\<lbrakk>continuous_on s f; x \<in> s\<rbrakk> \<Longrightarrow> \<bar>f x\<bar> \<le> normf f"
+    apply (simp add: normf_def)
+    apply (rule cSUP_upper, assumption)
+    by (simp add: bounded_imp_bdd_above compact compact_continuous_image compact_imp_bounded continuous_on_rabs)
+
+  lemma normf_least: "s \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> s \<Longrightarrow> \<bar>f x\<bar> \<le> M) \<Longrightarrow> normf f \<le> M"
+    by (simp add: normf_def cSUP_least)
+
+end
+
+lemma (in function_ring_on) one:
+  assumes U: "open U" and t0: "t0 \<in> s" "t0 \<in> U" and t1: "t1 \<in> s-U"
+    shows "\<exists>V. open V \<and> t0 \<in> V \<and> s \<inter> V \<subseteq> U \<and>
+               (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. f t > 1 - e))"
+proof -
+  have "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}" if t: "t \<in> s - U" for t
+  proof -
+    have "t \<noteq> t0" using t t0 by auto
+    then obtain g where g: "g \<in> R" "g t \<noteq> g t0"
+      using separable t0  by (metis Diff_subset subset_eq t)
+    define h where [abs_def]: "h x = g x - g t0" for x
+    have "h \<in> R"
+      unfolding h_def by (fast intro: g const diff)
+    then have hsq: "(\<lambda>w. (h w)\<^sup>2) \<in> R"
+      by (simp add: power2_eq_square mult)
+    have "h t \<noteq> h t0"
+      by (simp add: h_def g)
+    then have "h t \<noteq> 0"
+      by (simp add: h_def)
+    then have ht2: "0 < (h t)^2"
+      by simp
+    also have "... \<le> normf (\<lambda>w. (h w)\<^sup>2)"
+      using t normf_upper [where x=t] continuous [OF hsq] by force
+    finally have nfp: "0 < normf (\<lambda>w. (h w)\<^sup>2)" .
+    define p where [abs_def]: "p x = (1 / normf (\<lambda>w. (h w)\<^sup>2)) * (h x)^2" for x
+    have "p \<in> R"
+      unfolding p_def by (fast intro: hsq const mult)
+    moreover have "p t0 = 0"
+      by (simp add: p_def h_def)
+    moreover have "p t > 0"
+      using nfp ht2 by (simp add: p_def)
+    moreover have "\<And>x. x \<in> s \<Longrightarrow> p x \<in> {0..1}"
+      using nfp normf_upper [OF continuous [OF hsq] ] by (auto simp: p_def)
+    ultimately show "\<exists>pt \<in> R. pt t0 = 0 \<and> pt t > 0 \<and> pt ` s \<subseteq> {0..1}"
+      by auto
+  qed
+  then obtain pf where pf: "\<And>t. t \<in> s-U \<Longrightarrow> pf t \<in> R \<and> pf t t0 = 0 \<and> pf t t > 0"
+                   and pf01: "\<And>t. t \<in> s-U \<Longrightarrow> pf t ` s \<subseteq> {0..1}"
+    by metis
+  have com_sU: "compact (s-U)"
+    using compact closed_Int_compact U by (simp add: Diff_eq compact_Int_closed open_closed)
+  have "\<And>t. t \<in> s-U \<Longrightarrow> \<exists>A. open A \<and> A \<inter> s = {x\<in>s. 0 < pf t x}"
+    apply (rule open_Collect_positive)
+    by (metis pf continuous)
+  then obtain Uf where Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t) \<and> (Uf t) \<inter> s = {x\<in>s. 0 < pf t x}"
+    by metis
+  then have open_Uf: "\<And>t. t \<in> s-U \<Longrightarrow> open (Uf t)"
+    by blast
+  have tUft: "\<And>t. t \<in> s-U \<Longrightarrow> t \<in> Uf t"
+    using pf Uf by blast
+  then have *: "s-U \<subseteq> (\<Union>x \<in> s-U. Uf x)"
+    by blast
+  obtain subU where subU: "subU \<subseteq> s - U" "finite subU" "s - U \<subseteq> (\<Union>x \<in> subU. Uf x)"
+    by (blast intro: that open_Uf compactE_image [OF com_sU _ *])
+  then have [simp]: "subU \<noteq> {}"
+    using t1 by auto
+  then have cardp: "card subU > 0" using subU
+    by (simp add: card_gt_0_iff)
+  define p where [abs_def]: "p x = (1 / card subU) * (\<Sum>t \<in> subU. pf t x)" for x
+  have pR: "p \<in> R"
+    unfolding p_def using subU pf by (fast intro: pf const mult setsum)
+  have pt0 [simp]: "p t0 = 0"
+    using subU pf by (auto simp: p_def intro: setsum.neutral)
+  have pt_pos: "p t > 0" if t: "t \<in> s-U" for t
+  proof -
+    obtain i where i: "i \<in> subU" "t \<in> Uf i" using subU t by blast
+    show ?thesis
+      using subU i t
+      apply (clarsimp simp: p_def divide_simps)
+      apply (rule setsum_pos2 [OF \<open>finite subU\<close>])
+      using Uf t pf01 apply auto
+      apply (force elim!: subsetCE)
+      done
+  qed
+  have p01: "p x \<in> {0..1}" if t: "x \<in> s" for x
+  proof -
+    have "0 \<le> p x"
+      using subU cardp t
+      apply (simp add: p_def divide_simps setsum_nonneg)
+      apply (rule setsum_nonneg)
+      using pf01 by force
+    moreover have "p x \<le> 1"
+      using subU cardp t
+      apply (simp add: p_def divide_simps setsum_nonneg)
+      apply (rule setsum_bounded_above [where 'a=real and K=1, simplified])
+      using pf01 by force
+    ultimately show ?thesis
+      by auto
+  qed
+  have "compact (p ` (s-U))"
+    by (meson Diff_subset com_sU compact_continuous_image continuous continuous_on_subset pR)
+  then have "open (- (p ` (s-U)))"
+    by (simp add: compact_imp_closed open_Compl)
+  moreover have "0 \<in> - (p ` (s-U))"
+    by (metis (no_types) ComplI image_iff not_less_iff_gr_or_eq pt_pos)
+  ultimately obtain delta0 where delta0: "delta0 > 0" "ball 0 delta0 \<subseteq> - (p ` (s-U))"
+    by (auto simp: elim!: openE)
+  then have pt_delta: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> delta0"
+    by (force simp: ball_def dist_norm dest: p01)
+  define \<delta> where "\<delta> = delta0/2"
+  have "delta0 \<le> 1" using delta0 p01 [of t1] t1
+      by (force simp: ball_def dist_norm dest: p01)
+  with delta0 have \<delta>01: "0 < \<delta>" "\<delta> < 1"
+    by (auto simp: \<delta>_def)
+  have pt_\<delta>: "\<And>x. x \<in> s-U \<Longrightarrow> p x \<ge> \<delta>"
+    using pt_delta delta0 by (force simp: \<delta>_def)
+  have "\<exists>A. open A \<and> A \<inter> s = {x\<in>s. p x < \<delta>/2}"
+    by (rule open_Collect_less_Int [OF continuous [OF pR] continuous_on_const])
+  then obtain V where V: "open V" "V \<inter> s = {x\<in>s. p x < \<delta>/2}"
+    by blast
+  define k where "k = nat\<lfloor>1/\<delta>\<rfloor> + 1"
+  have "k>0"  by (simp add: k_def)
+  have "k-1 \<le> 1/\<delta>"
+    using \<delta>01 by (simp add: k_def)
+  with \<delta>01 have "k \<le> (1+\<delta>)/\<delta>"
+    by (auto simp: algebra_simps add_divide_distrib)
+  also have "... < 2/\<delta>"
+    using \<delta>01 by (auto simp: divide_simps)
+  finally have k2\<delta>: "k < 2/\<delta>" .
+  have "1/\<delta> < k"
+    using \<delta>01 unfolding k_def by linarith
+  with \<delta>01 k2\<delta> have k\<delta>: "1 < k*\<delta>" "k*\<delta> < 2"
+    by (auto simp: divide_simps)
+  define q where [abs_def]: "q n t = (1 - p t ^ n) ^ (k^n)" for n t
+  have qR: "q n \<in> R" for n
+    by (simp add: q_def const diff power pR)
+  have q01: "\<And>n t. t \<in> s \<Longrightarrow> q n t \<in> {0..1}"
+    using p01 by (simp add: q_def power_le_one algebra_simps)
+  have qt0 [simp]: "\<And>n. n>0 \<Longrightarrow> q n t0 = 1"
+    using t0 pf by (simp add: q_def power_0_left)
+  { fix t and n::nat
+    assume t: "t \<in> s \<inter> V"
+    with \<open>k>0\<close> V have "k * p t < k * \<delta> / 2"
+       by force
+    then have "1 - (k * \<delta> / 2)^n \<le> 1 - (k * p t)^n"
+      using  \<open>k>0\<close> p01 t by (simp add: power_mono)
+    also have "... \<le> q n t"
+      using Bernoulli_inequality [of "- ((p t)^n)" "k^n"]
+      apply (simp add: q_def)
+      by (metis IntE atLeastAtMost_iff p01 power_le_one power_mult_distrib t)
+    finally have "1 - (k * \<delta> / 2) ^ n \<le> q n t" .
+  } note limitV = this
+  { fix t and n::nat
+    assume t: "t \<in> s - U"
+    with \<open>k>0\<close> U have "k * \<delta> \<le> k * p t"
+      by (simp add: pt_\<delta>)
+    with k\<delta> have kpt: "1 < k * p t"
+      by (blast intro: less_le_trans)
+    have ptn_pos: "0 < p t ^ n"
+      using pt_pos [OF t] by simp
+    have ptn_le: "p t ^ n \<le> 1"
+      by (meson DiffE atLeastAtMost_iff p01 power_le_one t)
+    have "q n t = (1/(k^n * (p t)^n)) * (1 - p t ^ n) ^ (k^n) * k^n * (p t)^n"
+      using pt_pos [OF t] \<open>k>0\<close> by (simp add: q_def)
+    also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + k^n * (p t)^n)"
+      using pt_pos [OF t] \<open>k>0\<close>
+      apply simp
+      apply (simp only: times_divide_eq_right [symmetric])
+      apply (rule mult_left_mono [of "1::real", simplified])
+      apply (simp_all add: power_mult_distrib)
+      apply (rule zero_le_power)
+      using ptn_le by linarith
+    also have "... \<le> (1/(k * (p t))^n) * (1 - p t ^ n) ^ (k^n) * (1 + (p t)^n) ^ (k^n)"
+      apply (rule mult_left_mono [OF Bernoulli_inequality [of "p t ^ n" "k^n"]])
+      using \<open>k>0\<close> ptn_pos ptn_le
+      apply (auto simp: power_mult_distrib)
+      done
+    also have "... = (1/(k * (p t))^n) * (1 - p t ^ (2*n)) ^ (k^n)"
+      using pt_pos [OF t] \<open>k>0\<close>
+      by (simp add: algebra_simps power_mult power2_eq_square power_mult_distrib [symmetric])
+    also have "... \<le> (1/(k * (p t))^n) * 1"
+      apply (rule mult_left_mono [OF power_le_one])
+      using pt_pos \<open>k>0\<close> p01 power_le_one t apply auto
+      done
+    also have "... \<le> (1 / (k*\<delta>))^n"
+      using \<open>k>0\<close> \<delta>01  power_mono pt_\<delta> t
+      by (fastforce simp: field_simps)
+    finally have "q n t \<le> (1 / (real k * \<delta>)) ^ n " .
+  } note limitNonU = this
+  define NN
+    where "NN e = 1 + nat \<lceil>max (ln e / ln (real k * \<delta> / 2)) (- ln e / ln (real k * \<delta>))\<rceil>" for e
+  have NN: "of_nat (NN e) > ln e / ln (real k * \<delta> / 2)"  "of_nat (NN e) > - ln e / ln (real k * \<delta>)"
+              if "0<e" for e
+      unfolding NN_def  by linarith+
+  have NN1: "\<And>e. e>0 \<Longrightarrow> (k * \<delta> / 2)^NN e < e"
+    apply (subst Transcendental.ln_less_cancel_iff [symmetric])
+      prefer 3 apply (subst ln_realpow)
+    using \<open>k>0\<close> \<open>\<delta>>0\<close> NN  k\<delta>
+    apply (force simp add: field_simps)+
+    done
+  have NN0: "\<And>e. e>0 \<Longrightarrow> (1/(k*\<delta>))^NN e < e"
+    apply (subst Transcendental.ln_less_cancel_iff [symmetric])
+      prefer 3 apply (subst ln_realpow)
+    using \<open>k>0\<close> \<open>\<delta>>0\<close> NN k\<delta>
+    apply (force simp add: field_simps ln_div)+
+    done
+  { fix t and e::real
+    assume "e>0"
+    have "t \<in> s \<inter> V \<Longrightarrow> 1 - q (NN e) t < e" "t \<in> s - U \<Longrightarrow> q (NN e) t < e"
+    proof -
+      assume t: "t \<in> s \<inter> V"
+      show "1 - q (NN e) t < e"
+        by (metis add.commute diff_le_eq not_le limitV [OF t] less_le_trans [OF NN1 [OF \<open>e>0\<close>]])
+    next
+      assume t: "t \<in> s - U"
+      show "q (NN e) t < e"
+      using  limitNonU [OF t] less_le_trans [OF NN0 [OF \<open>e>0\<close>]] not_le by blast
+    qed
+  } then have "\<And>e. e > 0 \<Longrightarrow> \<exists>f\<in>R. f ` s \<subseteq> {0..1} \<and> (\<forall>t \<in> s \<inter> V. f t < e) \<and> (\<forall>t \<in> s - U. 1 - e < f t)"
+    using q01
+    by (rule_tac x="\<lambda>x. 1 - q (NN e) x" in bexI) (auto simp: algebra_simps intro: diff const qR)
+  moreover have t0V: "t0 \<in> V"  "s \<inter> V \<subseteq> U"
+    using pt_\<delta> t0 U V \<delta>01  by fastforce+
+  ultimately show ?thesis using V t0V
+    by blast
+qed
+
+text\<open>Non-trivial case, with @{term A} and @{term B} both non-empty\<close>
+lemma (in function_ring_on) two_special:
+  assumes A: "closed A" "A \<subseteq> s" "a \<in> A"
+      and B: "closed B" "B \<subseteq> s" "b \<in> B"
+      and disj: "A \<inter> B = {}"
+      and e: "0 < e" "e < 1"
+    shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
+proof -
+  { fix w
+    assume "w \<in> A"
+    then have "open ( - B)" "b \<in> s" "w \<notin> B" "w \<in> s"
+      using assms by auto
+    then have "\<exists>V. open V \<and> w \<in> V \<and> s \<inter> V \<subseteq> -B \<and>
+               (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> V. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
+      using one [of "-B" w b] assms \<open>w \<in> A\<close> by simp
+  }
+  then obtain Vf where Vf:
+         "\<And>w. w \<in> A \<Longrightarrow> open (Vf w) \<and> w \<in> Vf w \<and> s \<inter> Vf w \<subseteq> -B \<and>
+                         (\<forall>e>0. \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e))"
+    by metis
+  then have open_Vf: "\<And>w. w \<in> A \<Longrightarrow> open (Vf w)"
+    by blast
+  have tVft: "\<And>w. w \<in> A \<Longrightarrow> w \<in> Vf w"
+    using Vf by blast
+  then have setsum_max_0: "A \<subseteq> (\<Union>x \<in> A. Vf x)"
+    by blast
+  have com_A: "compact A" using A
+    by (metis compact compact_Int_closed inf.absorb_iff2)
+  obtain subA where subA: "subA \<subseteq> A" "finite subA" "A \<subseteq> (\<Union>x \<in> subA. Vf x)"
+    by (blast intro: that open_Vf compactE_image [OF com_A _ setsum_max_0])
+  then have [simp]: "subA \<noteq> {}"
+    using \<open>a \<in> A\<close> by auto
+  then have cardp: "card subA > 0" using subA
+    by (simp add: card_gt_0_iff)
+  have "\<And>w. w \<in> A \<Longrightarrow> \<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> s \<inter> Vf w. f x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. f x > 1 - e / card subA)"
+    using Vf e cardp by simp
+  then obtain ff where ff:
+         "\<And>w. w \<in> A \<Longrightarrow> ff w \<in> R \<and> ff w ` s \<subseteq> {0..1} \<and>
+                         (\<forall>x \<in> s \<inter> Vf w. ff w x < e / card subA) \<and> (\<forall>x \<in> s \<inter> B. ff w x > 1 - e / card subA)"
+    by metis
+  define pff where [abs_def]: "pff x = (\<Prod>w \<in> subA. ff w x)" for x
+  have pffR: "pff \<in> R"
+    unfolding pff_def using subA ff by (auto simp: intro: setprod)
+  moreover
+  have pff01: "pff x \<in> {0..1}" if t: "x \<in> s" for x
+  proof -
+    have "0 \<le> pff x"
+      using subA cardp t
+      apply (simp add: pff_def divide_simps setsum_nonneg)
+      apply (rule Groups_Big.linordered_semidom_class.setprod_nonneg)
+      using ff by fastforce
+    moreover have "pff x \<le> 1"
+      using subA cardp t
+      apply (simp add: pff_def divide_simps setsum_nonneg)
+      apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
+      using ff by fastforce
+    ultimately show ?thesis
+      by auto
+  qed
+  moreover
+  { fix v x
+    assume v: "v \<in> subA" and x: "x \<in> Vf v" "x \<in> s"
+    from subA v have "pff x = ff v x * (\<Prod>w \<in> subA - {v}. ff w x)"
+      unfolding pff_def  by (metis setprod.remove)
+    also have "... \<le> ff v x * 1"
+      apply (rule Rings.ordered_semiring_class.mult_left_mono)
+      apply (rule setprod_mono [where g = "\<lambda>x. 1", simplified])
+      using ff [THEN conjunct2, THEN conjunct1] v subA x
+      apply auto
+      apply (meson atLeastAtMost_iff contra_subsetD imageI)
+      apply (meson atLeastAtMost_iff contra_subsetD image_eqI)
+      using atLeastAtMost_iff by blast
+    also have "... < e / card subA"
+      using ff [THEN conjunct2, THEN conjunct2, THEN conjunct1] v subA x
+      by auto
+    also have "... \<le> e"
+      using cardp e by (simp add: divide_simps)
+    finally have "pff x < e" .
+  }
+  then have "\<And>x. x \<in> A \<Longrightarrow> pff x < e"
+    using A Vf subA by (metis UN_E contra_subsetD)
+  moreover
+  { fix x
+    assume x: "x \<in> B"
+    then have "x \<in> s"
+      using B by auto
+    have "1 - e \<le> (1 - e / card subA) ^ card subA"
+      using Bernoulli_inequality [of "-e / card subA" "card subA"] e cardp
+      by (auto simp: field_simps)
+    also have "... = (\<Prod>w \<in> subA. 1 - e / card subA)"
+      by (simp add: setprod_constant subA(2))
+    also have "... < pff x"
+      apply (simp add: pff_def)
+      apply (rule setprod_mono_strict [where f = "\<lambda>x. 1 - e / card subA", simplified])
+      apply (simp_all add: subA(2))
+      apply (intro ballI conjI)
+      using e apply (force simp: divide_simps)
+      using ff [THEN conjunct2, THEN conjunct2, THEN conjunct2] subA B x
+      apply blast
+      done
+    finally have "1 - e < pff x" .
+  }
+  ultimately
+  show ?thesis by blast
+qed
+
+lemma (in function_ring_on) two:
+  assumes A: "closed A" "A \<subseteq> s"
+      and B: "closed B" "B \<subseteq> s"
+      and disj: "A \<inter> B = {}"
+      and e: "0 < e" "e < 1"
+    shows "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A. f x < e) \<and> (\<forall>x \<in> B. f x > 1 - e)"
+proof (cases "A \<noteq> {} \<and> B \<noteq> {}")
+  case True then show ?thesis
+    apply (simp add: ex_in_conv [symmetric])
+    using assms
+    apply safe
+    apply (force simp add: intro!: two_special)
+    done
+next
+  case False with e show ?thesis
+    apply simp
+    apply (erule disjE)
+    apply (rule_tac [2] x="\<lambda>x. 0" in bexI)
+    apply (rule_tac x="\<lambda>x. 1" in bexI)
+    apply (auto simp: const)
+    done
+qed
+
+text\<open>The special case where @{term f} is non-negative and @{term"e<1/3"}\<close>
+lemma (in function_ring_on) Stone_Weierstrass_special:
+  assumes f: "continuous_on s f" and fpos: "\<And>x. x \<in> s \<Longrightarrow> f x \<ge> 0"
+      and e: "0 < e" "e < 1/3"
+  shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < 2*e"
+proof -
+  define n where "n = 1 + nat \<lceil>normf f / e\<rceil>"
+  define A where "A j = {x \<in> s. f x \<le> (j - 1/3)*e}" for j :: nat
+  define B where "B j = {x \<in> s. f x \<ge> (j + 1/3)*e}" for j :: nat
+  have ngt: "(n-1) * e \<ge> normf f" "n\<ge>1"
+    using e
+    apply (simp_all add: n_def field_simps of_nat_Suc)
+    by (metis real_nat_ceiling_ge mult.commute not_less pos_less_divide_eq)
+  then have ge_fx: "(n-1) * e \<ge> f x" if "x \<in> s" for x
+    using f normf_upper that by fastforce
+  { fix j
+    have A: "closed (A j)" "A j \<subseteq> s"
+      apply (simp_all add: A_def Collect_restrict)
+      apply (rule continuous_on_closed_Collect_le [OF f continuous_on_const])
+      apply (simp add: compact compact_imp_closed)
+      done
+    have B: "closed (B j)" "B j \<subseteq> s"
+      apply (simp_all add: B_def Collect_restrict)
+      apply (rule continuous_on_closed_Collect_le [OF continuous_on_const f])
+      apply (simp add: compact compact_imp_closed)
+      done
+    have disj: "(A j) \<inter> (B j) = {}"
+      using e by (auto simp: A_def B_def field_simps)
+    have "\<exists>f \<in> R. f ` s \<subseteq> {0..1} \<and> (\<forall>x \<in> A j. f x < e/n) \<and> (\<forall>x \<in> B j. f x > 1 - e/n)"
+      apply (rule two)
+      using e A B disj ngt
+      apply simp_all
+      done
+  }
+  then obtain xf where xfR: "\<And>j. xf j \<in> R" and xf01: "\<And>j. xf j ` s \<subseteq> {0..1}"
+                   and xfA: "\<And>x j. x \<in> A j \<Longrightarrow> xf j x < e/n"
+                   and xfB: "\<And>x j. x \<in> B j \<Longrightarrow> xf j x > 1 - e/n"
+    by metis
+  define g where [abs_def]: "g x = e * (\<Sum>i\<le>n. xf i x)" for x
+  have gR: "g \<in> R"
+    unfolding g_def by (fast intro: mult const setsum xfR)
+  have gge0: "\<And>x. x \<in> s \<Longrightarrow> g x \<ge> 0"
+    using e xf01 by (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
+  have A0: "A 0 = {}"
+    using fpos e by (fastforce simp: A_def)
+  have An: "A n = s"
+    using e ngt f normf_upper by (fastforce simp: A_def field_simps of_nat_diff)
+  have Asub: "A j \<subseteq> A i" if "i\<ge>j" for i j
+    using e that apply (clarsimp simp: A_def)
+    apply (erule order_trans, simp)
+    done
+  { fix t
+    assume t: "t \<in> s"
+    define j where "j = (LEAST j. t \<in> A j)"
+    have jn: "j \<le> n"
+      using t An by (simp add: Least_le j_def)
+    have Aj: "t \<in> A j"
+      using t An by (fastforce simp add: j_def intro: LeastI)
+    then have Ai: "t \<in> A i" if "i\<ge>j" for i
+      using Asub [OF that] by blast
+    then have fj1: "f t \<le> (j - 1/3)*e"
+      by (simp add: A_def)
+    then have Anj: "t \<notin> A i" if "i<j" for i
+      using  Aj  \<open>i<j\<close>
+      apply (simp add: j_def)
+      using not_less_Least by blast
+    have j1: "1 \<le> j"
+      using A0 Aj j_def not_less_eq_eq by (fastforce simp add: j_def)
+    then have Anj: "t \<notin> A (j-1)"
+      using Least_le by (fastforce simp add: j_def)
+    then have fj2: "(j - 4/3)*e < f t"
+      using j1 t  by (simp add: A_def of_nat_diff)
+    have ***: "xf i t \<le> e/n" if "i\<ge>j" for i
+      using xfA [OF Ai] that by (simp add: less_eq_real_def)
+    { fix i
+      assume "i+2 \<le> j"
+      then obtain d where "i+2+d = j"
+        using le_Suc_ex that by blast
+      then have "t \<in> B i"
+        using Anj e ge_fx [OF t] \<open>1 \<le> n\<close> fpos [OF t] t
+        apply (simp add: A_def B_def)
+        apply (clarsimp simp add: field_simps of_nat_diff not_le of_nat_Suc)
+        apply (rule order_trans [of _ "e * 2 + (e * (real d * 3) + e * (real i * 3))"])
+        apply auto
+        done
+      then have "xf i t > 1 - e/n"
+        by (rule xfB)
+    } note **** = this
+    have xf_le1: "\<And>i. xf i t \<le> 1"
+      using xf01 t by force
+    have "g t = e * (\<Sum>i<j. xf i t) + e * (\<Sum>i=j..n. xf i t)"
+      using j1 jn e
+      apply (simp add: g_def distrib_left [symmetric])
+      apply (subst setsum.union_disjoint [symmetric])
+      apply (auto simp: ivl_disj_un)
+      done
+    also have "... \<le> e*j + e * ((Suc n - j)*e/n)"
+      apply (rule add_mono)
+      apply (simp_all only: mult_le_cancel_left_pos e)
+      apply (rule setsum_bounded_above [OF xf_le1, where A = "lessThan j", simplified])
+      using setsum_bounded_above [of "{j..n}" "\<lambda>i. xf i t", OF ***]
+      apply simp
+      done
+    also have "... \<le> j*e + e*(n - j + 1)*e/n "
+      using \<open>1 \<le> n\<close> e  by (simp add: field_simps del: of_nat_Suc)
+    also have "... \<le> j*e + e*e"
+      using \<open>1 \<le> n\<close> e j1 by (simp add: field_simps del: of_nat_Suc)
+    also have "... < (j + 1/3)*e"
+      using e by (auto simp: field_simps)
+    finally have gj1: "g t < (j + 1 / 3) * e" .
+    have gj2: "(j - 4/3)*e < g t"
+    proof (cases "2 \<le> j")
+      case False
+      then have "j=1" using j1 by simp
+      with t gge0 e show ?thesis by force
+    next
+      case True
+      then have "(j - 4/3)*e < (j-1)*e - e^2"
+        using e by (auto simp: of_nat_diff algebra_simps power2_eq_square)
+      also have "... < (j-1)*e - ((j - 1)/n) * e^2"
+        using e True jn by (simp add: power2_eq_square field_simps)
+      also have "... = e * (j-1) * (1 - e/n)"
+        by (simp add: power2_eq_square field_simps)
+      also have "... \<le> e * (\<Sum>i\<le>j-2. xf i t)"
+        using e
+        apply simp
+        apply (rule order_trans [OF _ setsum_bounded_below [OF less_imp_le [OF ****]]])
+        using True
+        apply (simp_all add: of_nat_Suc of_nat_diff)
+        done
+      also have "... \<le> g t"
+        using jn e
+        using e xf01 t
+        apply (simp add: g_def zero_le_mult_iff image_subset_iff setsum_nonneg)
+        apply (rule Groups_Big.setsum_mono2, auto)
+        done
+      finally show ?thesis .
+    qed
+    have "\<bar>f t - g t\<bar> < 2 * e"
+      using fj1 fj2 gj1 gj2 by (simp add: abs_less_iff field_simps)
+  }
+  then show ?thesis
+    by (rule_tac x=g in bexI) (auto intro: gR)
+qed
+
+text\<open>The ``unpretentious'' formulation\<close>
+lemma (in function_ring_on) Stone_Weierstrass_basic:
+  assumes f: "continuous_on s f" and e: "e > 0"
+  shows "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>f x - g x\<bar> < e"
+proof -
+  have "\<exists>g \<in> R. \<forall>x\<in>s. \<bar>(f x + normf f) - g x\<bar> < 2 * min (e/2) (1/4)"
+    apply (rule Stone_Weierstrass_special)
+    apply (rule Limits.continuous_on_add [OF f Topological_Spaces.continuous_on_const])
+    using normf_upper [OF f] apply force
+    apply (simp add: e, linarith)
+    done
+  then obtain g where "g \<in> R" "\<forall>x\<in>s. \<bar>g x - (f x + normf f)\<bar> < e"
+    by force
+  then show ?thesis
+    apply (rule_tac x="\<lambda>x. g x - normf f" in bexI)
+    apply (auto simp: algebra_simps intro: diff const)
+    done
+qed
+
+
+theorem (in function_ring_on) Stone_Weierstrass:
+  assumes f: "continuous_on s f"
+  shows "\<exists>F\<in>UNIV \<rightarrow> R. LIM n sequentially. F n :> uniformly_on s f"
+proof -
+  { fix e::real
+    assume e: "0 < e"
+    then obtain N::nat where N: "0 < N" "0 < inverse N" "inverse N < e"
+      by (auto simp: real_arch_inverse [of e])
+    { fix n :: nat and x :: 'a and g :: "'a \<Rightarrow> real"
+      assume n: "N \<le> n"  "\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n)"
+      assume x: "x \<in> s"
+      have "\<not> real (Suc n) < inverse e"
+        using \<open>N \<le> n\<close> N using less_imp_inverse_less by force
+      then have "1 / (1 + real n) \<le> e"
+        using e by (simp add: field_simps of_nat_Suc)
+      then have "\<bar>f x - g x\<bar> < e"
+        using n(2) x by auto
+    } note * = this
+    have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>s. \<bar>f x - (SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + real n))) x\<bar> < e"
+      apply (rule eventually_sequentiallyI [of N])
+      apply (auto intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]] *)
+      done
+  } then
+  show ?thesis
+    apply (rule_tac x="\<lambda>n::nat. SOME g. g \<in> R \<and> (\<forall>x\<in>s. \<bar>f x - g x\<bar> < 1 / (1 + n))" in bexI)
+    prefer 2  apply (force intro: someI2_bex [OF Stone_Weierstrass_basic [OF f]])
+    unfolding uniform_limit_iff
+    apply (auto simp: dist_norm abs_minus_commute)
+    done
+qed
+
+text\<open>A HOL Light formulation\<close>
+corollary Stone_Weierstrass_HOL:
+  fixes R :: "('a::t2_space \<Rightarrow> real) set" and s :: "'a set"
+  assumes "compact s"  "\<And>c. P(\<lambda>x. c::real)"
+          "\<And>f. P f \<Longrightarrow> continuous_on s f"
+          "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x + g x)"  "\<And>f g. P(f) \<and> P(g) \<Longrightarrow> P(\<lambda>x. f x * g x)"
+          "\<And>x y. x \<in> s \<and> y \<in> s \<and> ~(x = y) \<Longrightarrow> \<exists>f. P(f) \<and> ~(f x = f y)"
+          "continuous_on s f"
+       "0 < e"
+    shows "\<exists>g. P(g) \<and> (\<forall>x \<in> s. \<bar>f x - g x\<bar> < e)"
+proof -
+  interpret PR: function_ring_on "Collect P"
+    apply unfold_locales
+    using assms
+    by auto
+  show ?thesis
+    using PR.Stone_Weierstrass_basic [OF \<open>continuous_on s f\<close> \<open>0 < e\<close>]
+    by blast
+qed
+
+
+subsection \<open>Polynomial functions\<close>
+
+inductive real_polynomial_function :: "('a::real_normed_vector \<Rightarrow> real) \<Rightarrow> bool" where
+    linear: "bounded_linear f \<Longrightarrow> real_polynomial_function f"
+  | const: "real_polynomial_function (\<lambda>x. c)"
+  | add:   "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x + g x)"
+  | mult:  "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x * g x)"
+
+declare real_polynomial_function.intros [intro]
+
+definition polynomial_function :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> bool"
+  where
+   "polynomial_function p \<equiv> (\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f o p))"
+
+lemma real_polynomial_function_eq: "real_polynomial_function p = polynomial_function p"
+unfolding polynomial_function_def
+proof
+  assume "real_polynomial_function p"
+  then show " \<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
+  proof (induction p rule: real_polynomial_function.induct)
+    case (linear h) then show ?case
+      by (auto simp: bounded_linear_compose real_polynomial_function.linear)
+  next
+    case (const h) then show ?case
+      by (simp add: real_polynomial_function.const)
+  next
+    case (add h) then show ?case
+      by (force simp add: bounded_linear_def linear_add real_polynomial_function.add)
+  next
+    case (mult h) then show ?case
+      by (force simp add: real_bounded_linear const real_polynomial_function.mult)
+  qed
+next
+  assume [rule_format, OF bounded_linear_ident]: "\<forall>f. bounded_linear f \<longrightarrow> real_polynomial_function (f \<circ> p)"
+  then show "real_polynomial_function p"
+    by (simp add: o_def)
+qed
+
+lemma polynomial_function_const [iff]: "polynomial_function (\<lambda>x. c)"
+  by (simp add: polynomial_function_def o_def const)
+
+lemma polynomial_function_bounded_linear:
+  "bounded_linear f \<Longrightarrow> polynomial_function f"
+  by (simp add: polynomial_function_def o_def bounded_linear_compose real_polynomial_function.linear)
+
+lemma polynomial_function_id [iff]: "polynomial_function(\<lambda>x. x)"
+  by (simp add: polynomial_function_bounded_linear)
+
+lemma polynomial_function_add [intro]:
+    "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x + g x)"
+  by (auto simp: polynomial_function_def bounded_linear_def linear_add real_polynomial_function.add o_def)
+
+lemma polynomial_function_mult [intro]:
+  assumes f: "polynomial_function f" and g: "polynomial_function g"
+    shows "polynomial_function (\<lambda>x. f x *\<^sub>R g x)"
+  using g
+  apply (auto simp: polynomial_function_def bounded_linear_def Real_Vector_Spaces.linear.scaleR  const real_polynomial_function.mult o_def)
+  apply (rule mult)
+  using f
+  apply (auto simp: real_polynomial_function_eq)
+  done
+
+lemma polynomial_function_cmul [intro]:
+  assumes f: "polynomial_function f"
+    shows "polynomial_function (\<lambda>x. c *\<^sub>R f x)"
+  by (rule polynomial_function_mult [OF polynomial_function_const f])
+
+lemma polynomial_function_minus [intro]:
+  assumes f: "polynomial_function f"
+    shows "polynomial_function (\<lambda>x. - f x)"
+  using polynomial_function_cmul [OF f, of "-1"] by simp
+
+lemma polynomial_function_diff [intro]:
+    "\<lbrakk>polynomial_function f; polynomial_function g\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. f x - g x)"
+  unfolding add_uminus_conv_diff [symmetric]
+  by (metis polynomial_function_add polynomial_function_minus)
+
+lemma polynomial_function_setsum [intro]:
+    "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> polynomial_function (\<lambda>x. setsum (f x) I)"
+by (induct I rule: finite_induct) auto
+
+lemma real_polynomial_function_minus [intro]:
+    "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. - f x)"
+  using polynomial_function_minus [of f]
+  by (simp add: real_polynomial_function_eq)
+
+lemma real_polynomial_function_diff [intro]:
+    "\<lbrakk>real_polynomial_function f; real_polynomial_function g\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. f x - g x)"
+  using polynomial_function_diff [of f]
+  by (simp add: real_polynomial_function_eq)
+
+lemma real_polynomial_function_setsum [intro]:
+    "\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> real_polynomial_function (\<lambda>x. f x i)\<rbrakk> \<Longrightarrow> real_polynomial_function (\<lambda>x. setsum (f x) I)"
+  using polynomial_function_setsum [of I f]
+  by (simp add: real_polynomial_function_eq)
+
+lemma real_polynomial_function_power [intro]:
+    "real_polynomial_function f \<Longrightarrow> real_polynomial_function (\<lambda>x. f x ^ n)"
+  by (induct n) (simp_all add: const mult)
+
+lemma real_polynomial_function_compose [intro]:
+  assumes f: "polynomial_function f" and g: "real_polynomial_function g"
+    shows "real_polynomial_function (g o f)"
+  using g
+  apply (induction g rule: real_polynomial_function.induct)
+  using f
+  apply (simp_all add: polynomial_function_def o_def const add mult)
+  done
+
+lemma polynomial_function_compose [intro]:
+  assumes f: "polynomial_function f" and g: "polynomial_function g"
+    shows "polynomial_function (g o f)"
+  using g real_polynomial_function_compose [OF f]
+  by (auto simp: polynomial_function_def o_def)
+
+lemma setsum_max_0:
+  fixes x::real (*in fact "'a::comm_ring_1"*)
+  shows "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..m. x^i * a i)"
+proof -
+  have "(\<Sum>i = 0..max m n. x^i * (if i \<le> m then a i else 0)) = (\<Sum>i = 0..max m n. (if i \<le> m then x^i * a i else 0))"
+    by (auto simp: algebra_simps intro: setsum.cong)
+  also have "... = (\<Sum>i = 0..m. (if i \<le> m then x^i * a i else 0))"
+    by (rule setsum.mono_neutral_right) auto
+  also have "... = (\<Sum>i = 0..m. x^i * a i)"
+    by (auto simp: algebra_simps intro: setsum.cong)
+  finally show ?thesis .
+qed
+
+lemma real_polynomial_function_imp_setsum:
+  assumes "real_polynomial_function f"
+    shows "\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i)"
+using assms
+proof (induct f)
+  case (linear f)
+  then show ?case
+    apply (clarsimp simp add: real_bounded_linear)
+    apply (rule_tac x="\<lambda>i. if i=0 then 0 else c" in exI)
+    apply (rule_tac x=1 in exI)
+    apply (simp add: mult_ac)
+    done
+next
+  case (const c)
+  show ?case
+    apply (rule_tac x="\<lambda>i. c" in exI)
+    apply (rule_tac x=0 in exI)
+    apply (auto simp: mult_ac of_nat_Suc)
+    done
+  case (add f1 f2)
+  then obtain a1 n1 a2 n2 where
+    "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
+    by auto
+  then show ?case
+    apply (rule_tac x="\<lambda>i. (if i \<le> n1 then a1 i else 0) + (if i \<le> n2 then a2 i else 0)" in exI)
+    apply (rule_tac x="max n1 n2" in exI)
+    using setsum_max_0 [where m=n1 and n=n2] setsum_max_0 [where m=n2 and n=n1]
+    apply (simp add: setsum.distrib algebra_simps max.commute)
+    done
+  case (mult f1 f2)
+  then obtain a1 n1 a2 n2 where
+    "f1 = (\<lambda>x. \<Sum>i = 0..n1. a1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. a2 i * x ^ i)"
+    by auto
+  then obtain b1 b2 where
+    "f1 = (\<lambda>x. \<Sum>i = 0..n1. b1 i * x ^ i)" "f2 = (\<lambda>x. \<Sum>i = 0..n2. b2 i * x ^ i)"
+    "b1 = (\<lambda>i. if i\<le>n1 then a1 i else 0)" "b2 = (\<lambda>i. if i\<le>n2 then a2 i else 0)"
+    by auto
+  then show ?case
+    apply (rule_tac x="\<lambda>i. \<Sum>k\<le>i. b1 k * b2 (i - k)" in exI)
+    apply (rule_tac x="n1+n2" in exI)
+    using polynomial_product [of n1 b1 n2 b2]
+    apply (simp add: Set_Interval.atLeast0AtMost)
+    done
+qed
+
+lemma real_polynomial_function_iff_setsum:
+     "real_polynomial_function f \<longleftrightarrow> (\<exists>a n::nat. f = (\<lambda>x. \<Sum>i=0..n. a i * x ^ i))"
+  apply (rule iffI)
+  apply (erule real_polynomial_function_imp_setsum)
+  apply (auto simp: linear mult const real_polynomial_function_power real_polynomial_function_setsum)
+  done
+
+lemma polynomial_function_iff_Basis_inner:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
+  shows "polynomial_function f \<longleftrightarrow> (\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. inner (f x) b))"
+        (is "?lhs = ?rhs")
+unfolding polynomial_function_def
+proof (intro iffI allI impI)
+  assume "\<forall>h. bounded_linear h \<longrightarrow> real_polynomial_function (h \<circ> f)"
+  then show ?rhs
+    by (force simp add: bounded_linear_inner_left o_def)
+next
+  fix h :: "'b \<Rightarrow> real"
+  assume rp: "\<forall>b\<in>Basis. real_polynomial_function (\<lambda>x. f x \<bullet> b)" and h: "bounded_linear h"
+  have "real_polynomial_function (h \<circ> (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b))"
+    apply (rule real_polynomial_function_compose [OF _  linear [OF h]])
+    using rp
+    apply (auto simp: real_polynomial_function_eq polynomial_function_mult)
+    done
+  then show "real_polynomial_function (h \<circ> f)"
+    by (simp add: euclidean_representation_setsum_fun)
+qed
+
+subsection \<open>Stone-Weierstrass theorem for polynomial functions\<close>
+
+text\<open>First, we need to show that they are continous, differentiable and separable.\<close>
+
+lemma continuous_real_polymonial_function:
+  assumes "real_polynomial_function f"
+    shows "continuous (at x) f"
+using assms
+by (induct f) (auto simp: linear_continuous_at)
+
+lemma continuous_polymonial_function:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
+  assumes "polynomial_function f"
+    shows "continuous (at x) f"
+  apply (rule euclidean_isCont)
+  using assms apply (simp add: polynomial_function_iff_Basis_inner)
+  apply (force dest: continuous_real_polymonial_function intro: isCont_scaleR)
+  done
+
+lemma continuous_on_polymonial_function:
+  fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::euclidean_space"
+  assumes "polynomial_function f"
+    shows "continuous_on s f"
+  using continuous_polymonial_function [OF assms] continuous_at_imp_continuous_on
+  by blast
+
+lemma has_real_derivative_polynomial_function:
+  assumes "real_polynomial_function p"
+    shows "\<exists>p'. real_polynomial_function p' \<and>
+                 (\<forall>x. (p has_real_derivative (p' x)) (at x))"
+using assms
+proof (induct p)
+  case (linear p)
+  then show ?case
+    by (force simp: real_bounded_linear const intro!: derivative_eq_intros)
+next
+  case (const c)
+  show ?case
+    by (rule_tac x="\<lambda>x. 0" in exI) auto
+  case (add f1 f2)
+  then obtain p1 p2 where
+    "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
+    "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
+    by auto
+  then show ?case
+    apply (rule_tac x="\<lambda>x. p1 x + p2 x" in exI)
+    apply (auto intro!: derivative_eq_intros)
+    done
+  case (mult f1 f2)
+  then obtain p1 p2 where
+    "real_polynomial_function p1" "\<And>x. (f1 has_real_derivative p1 x) (at x)"
+    "real_polynomial_function p2" "\<And>x. (f2 has_real_derivative p2 x) (at x)"
+    by auto
+  then show ?case
+    using mult
+    apply (rule_tac x="\<lambda>x. f1 x * p2 x + f2 x * p1 x" in exI)
+    apply (auto intro!: derivative_eq_intros)
+    done
+qed
+
+lemma has_vector_derivative_polynomial_function:
+  fixes p :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "polynomial_function p"
+    shows "\<exists>p'. polynomial_function p' \<and>
+                 (\<forall>x. (p has_vector_derivative (p' x)) (at x))"
+proof -
+  { fix b :: 'a
+    assume "b \<in> Basis"
+    then
+    obtain p' where p': "real_polynomial_function p'" and pd: "\<And>x. ((\<lambda>x. p x \<bullet> b) has_real_derivative p' x) (at x)"
+      using assms [unfolded polynomial_function_iff_Basis_inner, rule_format]  \<open>b \<in> Basis\<close>
+      has_real_derivative_polynomial_function
+      by blast
+    have "\<exists>q. polynomial_function q \<and> (\<forall>x. ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative q x) (at x))"
+      apply (rule_tac x="\<lambda>x. p' x *\<^sub>R b" in exI)
+      using \<open>b \<in> Basis\<close> p'
+      apply (simp add: polynomial_function_iff_Basis_inner inner_Basis)
+      apply (auto intro: derivative_eq_intros pd)
+      done
+  }
+  then obtain qf where qf:
+      "\<And>b. b \<in> Basis \<Longrightarrow> polynomial_function (qf b)"
+      "\<And>b x. b \<in> Basis \<Longrightarrow> ((\<lambda>u. (p u \<bullet> b) *\<^sub>R b) has_vector_derivative qf b x) (at x)"
+    by metis
+  show ?thesis
+    apply (subst euclidean_representation_setsum_fun [of p, symmetric])
+    apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. qf b x" in exI)
+    apply (auto intro: has_vector_derivative_setsum qf)
+    done
+qed
+
+lemma real_polynomial_function_separable:
+  fixes x :: "'a::euclidean_space"
+  assumes "x \<noteq> y" shows "\<exists>f. real_polynomial_function f \<and> f x \<noteq> f y"
+proof -
+  have "real_polynomial_function (\<lambda>u. \<Sum>b\<in>Basis. (inner (x-u) b)^2)"
+    apply (rule real_polynomial_function_setsum)
+    apply (auto simp: algebra_simps real_polynomial_function_power real_polynomial_function_diff
+                 const linear bounded_linear_inner_left)
+    done
+  then show ?thesis
+    apply (intro exI conjI, assumption)
+    using assms
+    apply (force simp add: euclidean_eq_iff [of x y] setsum_nonneg_eq_0_iff algebra_simps)
+    done
+qed
+
+lemma Stone_Weierstrass_real_polynomial_function:
+  fixes f :: "'a::euclidean_space \<Rightarrow> real"
+  assumes "compact s" "continuous_on s f" "0 < e"
+    shows "\<exists>g. real_polynomial_function g \<and> (\<forall>x \<in> s. \<bar>f x - g x\<bar> < e)"
+proof -
+  interpret PR: function_ring_on "Collect real_polynomial_function"
+    apply unfold_locales
+    using assms continuous_on_polymonial_function real_polynomial_function_eq
+    apply (auto intro: real_polynomial_function_separable)
+    done
+  show ?thesis
+    using PR.Stone_Weierstrass_basic [OF \<open>continuous_on s f\<close> \<open>0 < e\<close>]
+    by blast
+qed
+
+lemma Stone_Weierstrass_polynomial_function:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+  assumes s: "compact s"
+      and f: "continuous_on s f"
+      and e: "0 < e"
+    shows "\<exists>g. polynomial_function g \<and> (\<forall>x \<in> s. norm(f x - g x) < e)"
+proof -
+  { fix b :: 'b
+    assume "b \<in> Basis"
+    have "\<exists>p. real_polynomial_function p \<and> (\<forall>x \<in> s. \<bar>f x \<bullet> b - p x\<bar> < e / DIM('b))"
+      apply (rule exE [OF Stone_Weierstrass_real_polynomial_function [OF s _, of "\<lambda>x. f x \<bullet> b" "e / card Basis"]])
+      using e f
+      apply (auto simp: Euclidean_Space.DIM_positive intro: continuous_intros)
+      done
+  }
+  then obtain pf where pf:
+      "\<And>b. b \<in> Basis \<Longrightarrow> real_polynomial_function (pf b) \<and> (\<forall>x \<in> s. \<bar>f x \<bullet> b - pf b x\<bar> < e / DIM('b))"
+      apply (rule bchoice [rule_format, THEN exE])
+      apply assumption
+      apply (force simp add: intro: that)
+      done
+  have "polynomial_function (\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b)"
+    using pf
+    by (simp add: polynomial_function_setsum polynomial_function_mult real_polynomial_function_eq)
+  moreover
+  { fix x
+    assume "x \<in> s"
+    have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) \<le> (\<Sum>b\<in>Basis. norm ((f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b))"
+      by (rule norm_setsum)
+    also have "... < of_nat DIM('b) * (e / DIM('b))"
+      apply (rule setsum_bounded_above_strict)
+      apply (simp add: Real_Vector_Spaces.scaleR_diff_left [symmetric] pf \<open>x \<in> s\<close>)
+      apply (rule DIM_positive)
+      done
+    also have "... = e"
+      using DIM_positive by (simp add: field_simps)
+    finally have "norm (\<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b - pf b x *\<^sub>R b) < e" .
+  }
+  ultimately
+  show ?thesis
+    apply (subst euclidean_representation_setsum_fun [of f, symmetric])
+    apply (rule_tac x="\<lambda>x. \<Sum>b\<in>Basis. pf b x *\<^sub>R b" in exI)
+    apply (auto simp: setsum_subtractf [symmetric])
+    done
+qed
+
+
+subsection\<open>Polynomial functions as paths\<close>
+
+text\<open>One application is to pick a smooth approximation to a path,
+or just pick a smooth path anyway in an open connected set\<close>
+
+lemma path_polynomial_function:
+    fixes g  :: "real \<Rightarrow> 'b::euclidean_space"
+    shows "polynomial_function g \<Longrightarrow> path g"
+  by (simp add: path_def continuous_on_polymonial_function)
+
+lemma path_approx_polynomial_function:
+    fixes g :: "real \<Rightarrow> 'b::euclidean_space"
+    assumes "path g" "0 < e"
+    shows "\<exists>p. polynomial_function p \<and>
+                pathstart p = pathstart g \<and>
+                pathfinish p = pathfinish g \<and>
+                (\<forall>t \<in> {0..1}. norm(p t - g t) < e)"
+proof -
+  obtain q where poq: "polynomial_function q" and noq: "\<And>x. x \<in> {0..1} \<Longrightarrow> norm (g x - q x) < e/4"
+    using Stone_Weierstrass_polynomial_function [of "{0..1}" g "e/4"] assms
+    by (auto simp: path_def)
+  have pf: "polynomial_function (\<lambda>t. q t + (g 0 - q 0) + t *\<^sub>R (g 1 - q 1 - (g 0 - q 0)))"
+    by (force simp add: poq)
+  have *: "\<And>t. t \<in> {0..1} \<Longrightarrow> norm (((q t - g t) + (g 0 - q 0)) + (t *\<^sub>R (g 1 - q 1) + t *\<^sub>R (q 0 - g 0))) < (e/4 + e/4) + (e/4+e/4)"
+    apply (intro Real_Vector_Spaces.norm_add_less)
+    using noq
+    apply (auto simp: norm_minus_commute intro: le_less_trans [OF mult_left_le_one_le noq] simp del: less_divide_eq_numeral1)
+    done
+  show ?thesis
+    apply (intro exI conjI)
+    apply (rule pf)
+    using *
+    apply (auto simp add: pathstart_def pathfinish_def algebra_simps)
+    done
+qed
+
+lemma connected_open_polynomial_connected:
+  fixes s :: "'a::euclidean_space set"
+  assumes s: "open s" "connected s"
+      and "x \<in> s" "y \<in> s"
+    shows "\<exists>g. polynomial_function g \<and> path_image g \<subseteq> s \<and>
+               pathstart g = x \<and> pathfinish g = y"
+proof -
+  have "path_connected s" using assms
+    by (simp add: connected_open_path_connected)
+  with \<open>x \<in> s\<close> \<open>y \<in> s\<close> obtain p where p: "path p" "path_image p \<subseteq> s" "pathstart p = x" "pathfinish p = y"
+    by (force simp: path_connected_def)
+  have "\<exists>e. 0 < e \<and> (\<forall>x \<in> path_image p. ball x e \<subseteq> s)"
+  proof (cases "s = UNIV")
+    case True then show ?thesis
+      by (simp add: gt_ex)
+  next
+    case False
+    then have "- s \<noteq> {}" by blast
+    then show ?thesis
+      apply (rule_tac x="setdist (path_image p) (-s)" in exI)
+      using s p
+      apply (simp add: setdist_gt_0_compact_closed compact_path_image open_closed)
+      using setdist_le_dist [of _ "path_image p" _ "-s"]
+      by fastforce
+  qed
+  then obtain e where "0 < e"and eb: "\<And>x. x \<in> path_image p \<Longrightarrow> ball x e \<subseteq> s"
+    by auto
+  show ?thesis
+    using path_approx_polynomial_function [OF \<open>path p\<close> \<open>0 < e\<close>]
+    apply clarify
+    apply (intro exI conjI, assumption)
+    using p
+    apply (fastforce simp add: dist_norm path_image_def norm_minus_commute intro: eb [THEN subsetD])+
+    done
+qed
+
+hide_fact linear add mult const
+
+end
--- a/src/HOL/ROOT	Thu Aug 04 18:45:28 2016 +0200
+++ b/src/HOL/ROOT	Thu Aug 04 19:36:31 2016 +0200
@@ -719,12 +719,6 @@
 session "HOL-Multivariate_Analysis" (main) in Multivariate_Analysis = HOL +
   theories
     Multivariate_Analysis
-    Determinants
-    PolyRoots
-    Polytope
-    Complex_Analysis_Basics
-    Complex_Transcendental
-    Cauchy_Integral_Thm
   document_files
     "root.tex"