merged
authorwenzelm
Thu, 05 Jan 2017 22:38:06 +0100
changeset 64809 a0e1f64be67c
parent 64795 8e7db8df16a0 (diff)
parent 64808 81a5473e6d04 (current diff)
child 64810 05b29c8f0add
merged
--- a/src/HOL/Analysis/Analysis.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Analysis.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -12,6 +12,7 @@
   Weierstrass_Theorems
   Polytope
   Further_Topology
+  Arcwise_Connected
   Poly_Roots
   Conformal_Mappings
   Generalised_Binomial_Theorem
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Analysis/Arcwise_Connected.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -0,0 +1,2213 @@
+(*  Title:      HOL/Analysis/Arcwise_Connected.thy
+    Authors:    LC Paulson, based on material from HOL Light
+*)
+
+section \<open>Arcwise-connected sets\<close>
+
+theory Arcwise_Connected
+  imports Path_Connected Ordered_Euclidean_Space "~~/src/HOL/Number_Theory/Primes"
+
+begin
+
+subsection\<open>The Brouwer reduction theorem\<close>
+
+theorem Brouwer_reduction_theorem_gen:
+  fixes S :: "'a::euclidean_space set"
+  assumes "closed S" "\<phi> S"
+      and \<phi>: "\<And>F. \<lbrakk>\<And>n. closed(F n); \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>range F)"
+  obtains T where "T \<subseteq> S" "closed T" "\<phi> T" "\<And>U. \<lbrakk>U \<subseteq> S; closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
+proof -
+  obtain B :: "nat \<Rightarrow> 'a set"
+    where "inj B" "\<And>n. open(B n)" and open_cov: "\<And>S. open S \<Longrightarrow> \<exists>K. S = \<Union>(B ` K)"
+      by (metis Setcompr_eq_image that univ_second_countable_sequence)
+  define A where "A \<equiv> rec_nat S (\<lambda>n a. if \<exists>U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
+                                        then @U. U \<subseteq> a \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
+                                        else a)"
+  have [simp]: "A 0 = S"
+    by (simp add: A_def)
+  have ASuc: "A(Suc n) = (if \<exists>U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
+                          then @U. U \<subseteq> A n \<and> closed U \<and> \<phi> U \<and> U \<inter> (B n) = {}
+                          else A n)" for n
+    by (auto simp: A_def)
+  have sub: "\<And>n. A(Suc n) \<subseteq> A n"
+    by (auto simp: ASuc dest!: someI_ex)
+  have subS: "A n \<subseteq> S" for n
+    by (induction n) (use sub in auto)
+  have clo: "closed (A n) \<and> \<phi> (A n)" for n
+    by (induction n) (auto simp: assms ASuc dest!: someI_ex)
+  show ?thesis
+  proof
+    show "\<Inter>range A \<subseteq> S"
+      using \<open>\<And>n. A n \<subseteq> S\<close> by blast
+    show "closed (INTER UNIV A)"
+      using clo by blast
+    show "\<phi> (INTER UNIV A)"
+      by (simp add: clo \<phi> sub)
+    show "\<not> U \<subset> INTER UNIV A" if "U \<subseteq> S" "closed U" "\<phi> U" for U
+    proof -
+      have "\<exists>y. x \<notin> A y" if "x \<notin> U" and Usub: "U \<subseteq> (\<Inter>x. A x)" for x
+      proof -
+        obtain e where "e > 0" and e: "ball x e \<subseteq> -U"
+          using \<open>closed U\<close> \<open>x \<notin> U\<close> openE [of "-U"] by blast
+        moreover obtain K where K: "ball x e = UNION K B"
+          using open_cov [of "ball x e"] by auto
+        ultimately have "UNION K B \<subseteq> -U"
+          by blast
+        have "K \<noteq> {}"
+          using \<open>0 < e\<close> \<open>ball x e = UNION K B\<close> by auto
+        then obtain n where "n \<in> K" "x \<in> B n"
+          by (metis K UN_E \<open>0 < e\<close> centre_in_ball)
+        then have "U \<inter> B n = {}"
+          using K e by auto
+        show ?thesis
+        proof (cases "\<exists>U\<subseteq>A n. closed U \<and> \<phi> U \<and> U \<inter> B n = {}")
+          case True
+          then show ?thesis
+            apply (rule_tac x="Suc n" in exI)
+            apply (simp add: ASuc)
+            apply (erule someI2_ex)
+            using \<open>x \<in> B n\<close> by blast
+        next
+          case False
+          then show ?thesis
+            by (meson Inf_lower Usub \<open>U \<inter> B n = {}\<close> \<open>\<phi> U\<close> \<open>closed U\<close> range_eqI subset_trans)
+        qed
+      qed
+      with that show ?thesis
+        by (meson Inter_iff psubsetE rangeI subsetI)
+    qed
+  qed
+qed
+
+corollary Brouwer_reduction_theorem:
+  fixes S :: "'a::euclidean_space set"
+  assumes "compact S" "\<phi> S" "S \<noteq> {}"
+      and \<phi>: "\<And>F. \<lbrakk>\<And>n. compact(F n); \<And>n. F n \<noteq> {}; \<And>n. \<phi>(F n); \<And>n. F(Suc n) \<subseteq> F n\<rbrakk> \<Longrightarrow> \<phi>(\<Inter>range F)"
+  obtains T where "T \<subseteq> S" "compact T" "T \<noteq> {}" "\<phi> T"
+                  "\<And>U. \<lbrakk>U \<subseteq> S; closed U; U \<noteq> {}; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
+proof (rule Brouwer_reduction_theorem_gen [of S "\<lambda>T. T \<noteq> {} \<and> T \<subseteq> S \<and> \<phi> T"])
+  fix F
+  assume cloF: "\<And>n. closed (F n)"
+     and F: "\<And>n. F n \<noteq> {} \<and> F n \<subseteq> S \<and> \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n"
+  show "INTER UNIV F \<noteq> {} \<and> INTER UNIV F \<subseteq> S \<and> \<phi> (INTER UNIV F)"
+  proof (intro conjI)
+    show "INTER UNIV F \<noteq> {}"
+      apply (rule compact_nest)
+      apply (meson F cloF \<open>compact S\<close> seq_compact_closed_subset seq_compact_eq_compact)
+       apply (simp add: F)
+      by (meson Fsub lift_Suc_antimono_le)
+    show " INTER UNIV F \<subseteq> S"
+      using F by blast
+    show "\<phi> (INTER UNIV F)"
+      by (metis F Fsub \<phi> \<open>compact S\<close> cloF closed_Int_compact inf.orderE)
+  qed
+next
+  show "S \<noteq> {} \<and> S \<subseteq> S \<and> \<phi> S"
+    by (simp add: assms)
+qed (meson assms compact_imp_closed seq_compact_closed_subset seq_compact_eq_compact)+
+
+
+subsection\<open>Arcwise Connections\<close>
+
+subsection\<open>Density of points with dyadic rational coordinates.\<close>
+
+proposition closure_dyadic_rationals:
+    "closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
+                   { \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i }) = UNIV"
+proof -
+  have "x \<in> closure (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. {\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i})" for x::'a
+  proof (clarsimp simp: closure_approachable)
+    fix e::real
+    assume "e > 0"
+    then obtain k where k: "(1/2)^k < e/DIM('a)"
+      by (meson DIM_positive divide_less_eq_1_pos of_nat_0_less_iff one_less_numeral_iff real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral)
+    have "dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x =
+          dist (\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) (\<Sum>i\<in>Basis. (x \<bullet> i) *\<^sub>R i)"
+      by (simp add: euclidean_representation)
+    also have "... = norm ((\<Sum>i\<in>Basis. (real_of_int \<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i))"
+      by (simp add: dist_norm sum_subtractf)
+    also have "... \<le> DIM('a)*((1/2)^k)"
+    proof (rule sum_norm_bound, simp add: algebra_simps)
+      fix i::'a
+      assume "i \<in> Basis"
+      then have "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) =
+                 \<bar>real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k - x \<bullet> i\<bar>"
+        by (simp add: scaleR_left_diff_distrib [symmetric])
+      also have "... \<le> (1/2) ^ k"
+        by (simp add: divide_simps) linarith
+      finally show "norm ((real_of_int \<lfloor>x \<bullet> i*2^k\<rfloor> / 2^k) *\<^sub>R i - (x \<bullet> i) *\<^sub>R i) \<le> (1/2) ^ k" .
+    qed
+    also have "... < DIM('a)*(e/DIM('a))"
+      using DIM_positive k linordered_comm_semiring_strict_class.comm_mult_strict_left_mono of_nat_0_less_iff by blast
+    also have "... = e"
+      by simp
+    finally have "dist (\<Sum>i\<in>Basis. (\<lfloor>2^k*(x \<bullet> i)\<rfloor> / 2^k) *\<^sub>R i) x < e" .
+    then
+    show "\<exists>k. \<exists>f \<in> Basis \<rightarrow> \<int>. dist (\<Sum>b\<in>Basis. (f b / 2^k) *\<^sub>R b) x < e"
+      apply (rule_tac x=k in exI)
+      apply (rule_tac x="\<lambda>i. of_int (floor (2^k*(x \<bullet> i)))" in bexI)
+       apply auto
+      done
+  qed
+  then show ?thesis by auto
+qed
+
+corollary closure_rational_coordinates:
+    "closure (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i }) = UNIV"
+proof -
+  have *: "(\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>. { \<Sum>i::'a \<in> Basis. (f i / 2^k) *\<^sub>R i })
+           \<subseteq> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i \<in> Basis. f i *\<^sub>R i })"
+  proof clarsimp
+    fix k and f :: "'a \<Rightarrow> real"
+    assume f: "f \<in> Basis \<rightarrow> \<int>"
+    show "\<exists>x \<in> Basis \<rightarrow> \<rat>. (\<Sum>i \<in> Basis. (f i / 2^k) *\<^sub>R i) = (\<Sum>i \<in> Basis. x i *\<^sub>R i)"
+      apply (rule_tac x="\<lambda>i. f i / 2^k" in bexI)
+      using Ints_subset_Rats f by auto
+  qed
+  show ?thesis
+    using closure_dyadic_rationals closure_mono [OF *] by blast
+qed
+
+lemma closure_dyadic_rationals_in_convex_set:
+   "\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
+        \<Longrightarrow> closure(S \<inter>
+                    (\<Union>k. \<Union>f \<in> Basis \<rightarrow> \<int>.
+                   { \<Sum>i :: 'a :: euclidean_space \<in> Basis. (f i / 2^k) *\<^sub>R i })) =
+            closure S"
+  by (simp add: closure_dyadic_rationals closure_convex_Int_superset)
+
+lemma closure_rationals_in_convex_set:
+   "\<lbrakk>convex S; interior S \<noteq> {}\<rbrakk>
+    \<Longrightarrow> closure(S \<inter> (\<Union>f \<in> Basis \<rightarrow> \<rat>. { \<Sum>i :: 'a :: euclidean_space \<in> Basis. f i *\<^sub>R i })) =
+        closure S"
+  by (simp add: closure_rational_coordinates closure_convex_Int_superset)
+
+
+text\<open> Every path between distinct points contains an arc, and hence
+path connection is equivalent to arcwise connection for distinct points.
+The proof is based on Whyburn's "Topological Analysis".\<close>
+
+lemma closure_dyadic_rationals_in_convex_set_pos_1:
+  fixes S :: "real set"
+  assumes "convex S" and intnz: "interior S \<noteq> {}" and pos: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> x"
+    shows "closure(S \<inter> (\<Union>k m. {of_nat m / 2^k})) = closure S"
+proof -
+  have "\<exists>m. f 1/2^k = real m / 2^k" if "(f 1) / 2^k \<in> S" "f 1 \<in> \<int>" for k and f :: "real \<Rightarrow> real"
+    using that by (force simp: Ints_def zero_le_divide_iff power_le_zero_eq dest: pos zero_le_imp_eq_int)
+  then have "S \<inter> (\<Union>k m. {real m / 2^k}) = S \<inter>
+             (\<Union>k. \<Union>f\<in>Basis \<rightarrow> \<int>. {\<Sum>i\<in>Basis. (f i / 2^k) *\<^sub>R i})"
+    by force
+  then show ?thesis
+    using closure_dyadic_rationals_in_convex_set [OF \<open>convex S\<close> intnz] by simp
+qed
+
+
+definition dyadics :: "'a::field_char_0 set" where "dyadics \<equiv> \<Union>k m. {of_nat m / 2^k}"
+
+lemma real_in_dyadics [simp]: "real m \<in> dyadics"
+  apply (simp add: dyadics_def)
+  by (metis divide_numeral_1 numeral_One power_0)
+
+lemma nat_neq_4k1: "of_nat m \<noteq> (4 * of_nat k + 1) / (2 * 2^n :: 'a::field_char_0)"
+proof
+  assume "of_nat m = (4 * of_nat k + 1) / (2 * 2^n :: 'a)"
+  then have "of_nat (m * (2 * 2^n)) = (of_nat (Suc (4 * k)) :: 'a)"
+    by (simp add: divide_simps)
+  then have "m * (2 * 2^n) = Suc (4 * k)"
+    using of_nat_eq_iff by blast
+  then have "odd (m * (2 * 2^n))"
+    by simp
+  then show False
+    by simp
+qed
+
+lemma nat_neq_4k3: "of_nat m \<noteq> (4 * of_nat k + 3) / (2 * 2^n :: 'a::field_char_0)"
+proof
+  assume "of_nat m = (4 * of_nat k + 3) / (2 * 2^n :: 'a)"
+  then have "of_nat (m * (2 * 2^n)) = (of_nat (4 * k + 3) :: 'a)"
+    by (simp add: divide_simps)
+  then have "m * (2 * 2^n) = (4 * k) + 3"
+    using of_nat_eq_iff by blast
+  then have "odd (m * (2 * 2^n))"
+    by simp
+  then show False
+    by simp
+qed
+
+lemma iff_4k:
+  assumes "r = real k" "odd k"
+    shows "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n') \<longleftrightarrow> m=m' \<and> n=n'"
+proof -
+  { assume "(4 * real m + r) / (2 * 2^n) = (4 * real m' + r) / (2 * 2 ^ n')"
+    then have "real ((4 * m + k) * (2 * 2 ^ n')) = real ((4 * m' + k) * (2 * 2^n))"
+      using assms by (auto simp: field_simps)
+    then have "(4 * m + k) * (2 * 2 ^ n') = (4 * m' + k) * (2 * 2^n)"
+      using of_nat_eq_iff by blast
+    then have "(4 * m + k) * (2 ^ n') = (4 * m' + k) * (2^n)"
+      by linarith
+    then obtain "4*m + k = 4*m' + k" "n=n'"
+      apply (rule prime_power_cancel2 [OF two_is_prime_nat])
+      using assms by auto
+    then have "m=m'" "n=n'"
+      by auto
+  }
+  then show ?thesis by blast
+qed
+
+lemma neq_4k1_k43: "(4 * real m + 1) / (2 * 2^n) \<noteq> (4 * real m' + 3) / (2 * 2 ^ n')"
+proof
+  assume "(4 * real m + 1) / (2 * 2^n) = (4 * real m' + 3) / (2 * 2 ^ n')"
+  then have "real (Suc (4 * m) * (2 * 2 ^ n')) = real ((4 * m' + 3) * (2 * 2^n))"
+    by (auto simp: field_simps)
+  then have "Suc (4 * m) * (2 * 2 ^ n') = (4 * m' + 3) * (2 * 2^n)"
+    using of_nat_eq_iff by blast
+  then have "Suc (4 * m) * (2 ^ n') = (4 * m' + 3) * (2^n)"
+    by linarith
+  then have "Suc (4 * m) = (4 * m' + 3)"
+    by (rule prime_power_cancel2 [OF two_is_prime_nat]) auto
+  then have "1 + 2 * m' = 2 * m"
+    using \<open>Suc (4 * m) = 4 * m' + 3\<close> by linarith
+  then show False
+    using even_Suc by presburger
+qed
+
+lemma dyadic_413_cases:
+  obtains "(of_nat m::'a::field_char_0) / 2^k \<in> Nats"
+  | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 1) / 2^Suc k'"
+  | m' k' where "k' < k" "(of_nat m:: 'a) / 2^k = of_nat (4*m' + 3) / 2^Suc k'"
+proof (cases "m>0")
+  case False
+  then have "m=0" by simp
+  with that show ?thesis by auto
+next
+  case True
+  obtain k' m' where m': "odd m'" and k': "m = m' * 2^k'"
+    using prime_power_canonical [OF two_is_prime_nat True] by blast
+  then obtain q r where q: "m' = 4*q + r" and r: "r < 4"
+    by (metis not_add_less2 split_div zero_neq_numeral)
+  show ?thesis
+  proof (cases "k \<le> k'")
+    case True
+    have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)"
+      using k' by (simp add: field_simps)
+    also have "... = (of_nat m'::'a) * 2 ^ (k'-k)"
+      using k' True by (simp add: power_diff)
+    also have "... \<in> \<nat>"
+      by (metis Nats_mult of_nat_in_Nats of_nat_numeral of_nat_power)
+    finally show ?thesis by (auto simp: that)
+  next
+    case False
+    then obtain kd where kd: "Suc kd = k - k'"
+      using Suc_diff_Suc not_less by blast
+    have "(of_nat m:: 'a) / 2^k = of_nat m' * (2 ^ k' / 2^k)"
+      using k' by (simp add: field_simps)
+    also have "... = (of_nat m'::'a) / 2 ^ (k-k')"
+      using k' False by (simp add: power_diff)
+    also have "... = ((of_nat r + 4 * of_nat q)::'a) / 2 ^ (k-k')"
+      using q by force
+    finally have meq: "(of_nat m:: 'a) / 2^k = (of_nat r + 4 * of_nat q) / 2 ^ (k - k')" .
+    have "r \<noteq> 0" "r \<noteq> 2"
+      using q m' by presburger+
+    with r consider "r = 1" | "r = 3"
+      by linarith
+    then show ?thesis
+    proof cases
+      assume "r = 1"
+      with meq kd that(2) [of kd q] show ?thesis
+        by simp
+    next
+      assume "r = 3"
+      with meq kd that(3) [of kd q]  show ?thesis
+        by simp
+    qed
+  qed
+qed
+
+
+lemma dyadics_iff:
+   "(dyadics :: 'a::field_char_0 set) =
+    Nats \<union> (\<Union>k m. {of_nat (4*m + 1) / 2^Suc k}) \<union> (\<Union>k m. {of_nat (4*m + 3) / 2^Suc k})"
+           (is "_ = ?rhs")
+proof
+  show "dyadics \<subseteq> ?rhs"
+    unfolding dyadics_def
+    apply clarify
+    apply (rule dyadic_413_cases, force+)
+    done
+next
+  show "?rhs \<subseteq> dyadics"
+    apply (clarsimp simp: dyadics_def Nats_def simp del: power_Suc)
+    apply (intro conjI subsetI)
+    apply (auto simp del: power_Suc)
+      apply (metis divide_numeral_1 numeral_One power_0)
+     apply (metis of_nat_Suc of_nat_mult of_nat_numeral)
+    by (metis of_nat_add of_nat_mult of_nat_numeral)
+qed
+
+
+function (domintros) dyad_rec :: "[nat \<Rightarrow> 'a, 'a\<Rightarrow>'a, 'a\<Rightarrow>'a, real] \<Rightarrow> 'a" where
+    "dyad_rec b l r (real m) = b m"
+  | "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
+  | "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
+  | "x \<notin> dyadics \<Longrightarrow> dyad_rec b l r x = undefined"
+  using iff_4k [of _ 1] iff_4k [of _ 3]
+         apply (simp_all add: nat_neq_4k1 nat_neq_4k3 neq_4k1_k43, atomize_elim)
+     apply (fastforce simp add: dyadics_iff Nats_def field_simps)+
+  done
+
+lemma dyadics_levels: "dyadics = (\<Union>K. \<Union>k<K. \<Union> m. {of_nat m / 2^k})"
+  unfolding dyadics_def by auto
+
+lemma dyad_rec_level_termination:
+  assumes "k < K"
+  shows "dyad_rec_dom(b, l, r, real m / 2^k)"
+  using assms
+proof (induction K arbitrary: k m)
+  case 0
+  then show ?case by auto
+next
+  case (Suc K)
+  then consider "k = K" | "k < K"
+    using less_antisym by blast
+  then show ?case
+  proof cases
+    assume "k = K"
+    show ?case
+    proof (rule dyadic_413_cases [of m k, where 'a=real])
+      show "real m / 2^k \<in> \<nat> \<Longrightarrow> dyad_rec_dom (b, l, r, real m / 2^k)"
+        by (force simp: Nats_def nat_neq_4k1 nat_neq_4k3 intro: dyad_rec.domintros)
+      show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 1) / 2^Suc k'" for m' k'
+      proof -
+        have "dyad_rec_dom (b, l, r, (4 * real m' + 1) / 2^Suc k')"
+        proof (rule dyad_rec.domintros)
+          fix m n
+          assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)"
+          then have "m' = m" "k' = n" using iff_4k [of _ 1]
+            by auto
+          have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')"
+            using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast
+          then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)"
+            using \<open>k' = n\<close> by (auto simp: algebra_simps)
+        next
+          fix m n
+          assume "(4 * real m' + 1) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)"
+          then have "False"
+            by (metis neq_4k1_k43)
+          then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" ..
+        qed
+        then show ?case by (simp add: eq add_ac)
+      qed
+      show ?case if "k' < k" and eq: "real m / 2^k = real (4 * m' + 3) / 2^Suc k'" for m' k'
+      proof -
+        have "dyad_rec_dom (b, l, r, (4 * real m' + 3) / 2^Suc k')"
+        proof (rule dyad_rec.domintros)
+          fix m n
+          assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 1) / (2 * 2^n)"
+          then have "False"
+            by (metis neq_4k1_k43)
+          then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)" ..
+        next
+          fix m n
+          assume "(4 * real m' + 3) / (2 * 2 ^ k') = (4 * real m + 3) / (2 * 2^n)"
+          then have "m' = m" "k' = n" using iff_4k [of _ 3]
+            by auto
+          have "dyad_rec_dom (b, l, r, real (2 * m + 1) / 2 ^ k')"
+            using Suc.IH \<open>k = K\<close> \<open>k' < k\<close> by blast
+          then show "dyad_rec_dom (b, l, r, (2 * real m + 1) / 2^n)"
+            using \<open>k' = n\<close> by (auto simp: algebra_simps)
+        qed
+        then show ?case by (simp add: eq add_ac)
+      qed
+    qed
+  next
+    assume "k < K"
+    then show ?case
+      using Suc.IH by blast
+  qed
+qed
+
+
+lemma dyad_rec_termination: "x \<in> dyadics \<Longrightarrow> dyad_rec_dom(b,l,r,x)"
+  by (auto simp: dyadics_levels intro: dyad_rec_level_termination)
+
+lemma dyad_rec_of_nat [simp]: "dyad_rec b l r (real m) = b m"
+  by (simp add: dyad_rec.psimps dyad_rec_termination)
+
+lemma dyad_rec_41 [simp]: "dyad_rec b l r ((4 * real m + 1) / 2 ^ (Suc n)) = l (dyad_rec b l r ((2*m + 1) / 2^n))"
+  apply (rule dyad_rec.psimps)
+  by (metis dyad_rec_level_termination lessI add.commute of_nat_Suc of_nat_mult of_nat_numeral)
+
+
+lemma dyad_rec_43 [simp]: "dyad_rec b l r ((4 * real m + 3) / 2 ^ (Suc n)) = r (dyad_rec b l r ((2*m + 1) / 2^n))"
+  apply (rule dyad_rec.psimps)
+  by (metis dyad_rec_level_termination lessI of_nat_add of_nat_mult of_nat_numeral)
+
+lemma dyad_rec_41_times2:
+  assumes "n > 0"
+    shows "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
+proof -
+  obtain n' where n': "n = Suc n'"
+    using assms not0_implies_Suc by blast
+  have "dyad_rec b l r (2 * ((4 * real m + 1) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 1)) / (2 * 2^n))"
+    by auto
+  also have "... = dyad_rec b l r ((4 * real m + 1) / 2^n)"
+    by (subst mult_divide_mult_cancel_left) auto
+  also have "... = l (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))"
+    by (simp add: add.commute [of 1] n' del: power_Suc)
+  also have "... = l (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))"
+    by (subst mult_divide_mult_cancel_left) auto
+  also have "... = l (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
+    by (simp add: add.commute n')
+  finally show ?thesis .
+qed
+
+lemma dyad_rec_43_times2:
+  assumes "n > 0"
+    shows "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
+proof -
+  obtain n' where n': "n = Suc n'"
+    using assms not0_implies_Suc by blast
+  have "dyad_rec b l r (2 * ((4 * real m + 3) / 2^Suc n)) = dyad_rec b l r ((2 * (4 * real m + 3)) / (2 * 2^n))"
+    by auto
+  also have "... = dyad_rec b l r ((4 * real m + 3) / 2^n)"
+    by (subst mult_divide_mult_cancel_left) auto
+  also have "... = r (dyad_rec b l r ((2 * real m + 1) / 2 ^ n'))"
+    by (simp add: n' del: power_Suc)
+  also have "... = r (dyad_rec b l r ((2 * (2 * real m + 1)) / (2 * 2 ^ n')))"
+    by (subst mult_divide_mult_cancel_left) auto
+  also have "... = r (dyad_rec b l r (2 * (2 * real m + 1) / 2^n))"
+    by (simp add: n')
+  finally show ?thesis .
+qed
+
+definition dyad_rec2
+    where "dyad_rec2 u v lc rc x =
+             dyad_rec (\<lambda>z. (u,v)) (\<lambda>(a,b). (a, lc a b (midpoint a b))) (\<lambda>(a,b). (rc a b (midpoint a b), b)) (2*x)"
+
+abbreviation leftrec where "leftrec u v lc rc x \<equiv> fst (dyad_rec2 u v lc rc x)"
+abbreviation rightrec where "rightrec u v lc rc x \<equiv> snd (dyad_rec2 u v lc rc x)"
+
+lemma leftrec_base: "leftrec u v lc rc (real m / 2) = u"
+  by (simp add: dyad_rec2_def)
+
+lemma leftrec_41: "n > 0 \<Longrightarrow> leftrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) = leftrec u v lc rc ((2 * real m + 1) / 2^n)"
+  apply (simp only: dyad_rec2_def dyad_rec_41_times2)
+  apply (simp add: case_prod_beta)
+  done
+
+lemma leftrec_43: "n > 0 \<Longrightarrow>
+             leftrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) =
+                 rc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
+                 (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
+  apply (simp only: dyad_rec2_def dyad_rec_43_times2)
+  apply (simp add: case_prod_beta)
+  done
+
+lemma rightrec_base: "rightrec u v lc rc (real m / 2) = v"
+  by (simp add: dyad_rec2_def)
+
+lemma rightrec_41: "n > 0 \<Longrightarrow>
+             rightrec u v lc rc ((4 * real m + 1) / 2 ^ (Suc n)) =
+                 lc (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n))
+                 (midpoint (leftrec u v lc rc ((2 * real m + 1) / 2^n)) (rightrec u v lc rc ((2 * real m + 1) / 2^n)))"
+  apply (simp only: dyad_rec2_def dyad_rec_41_times2)
+  apply (simp add: case_prod_beta)
+  done
+
+lemma rightrec_43: "n > 0 \<Longrightarrow> rightrec u v lc rc ((4 * real m + 3) / 2 ^ (Suc n)) = rightrec u v lc rc ((2 * real m + 1) / 2^n)"
+  apply (simp only: dyad_rec2_def dyad_rec_43_times2)
+  apply (simp add: case_prod_beta)
+  done
+
+lemma dyadics_in_open_unit_interval:
+   "{0<..<1} \<inter> (\<Union>k m. {real m / 2^k}) = (\<Union>k. \<Union>m \<in> {0<..<2^k}. {real m / 2^k})"
+  by (auto simp: divide_simps)
+
+
+
+theorem homeomorphic_monotone_image_interval:
+  fixes f :: "real \<Rightarrow> 'a::{real_normed_vector,complete_space}"
+  assumes cont_f: "continuous_on {0..1} f"
+      and conn: "\<And>y. connected ({0..1} \<inter> f -` {y})"
+      and f_1not0: "f 1 \<noteq> f 0"
+    shows "(f ` {0..1}) homeomorphic {0..1::real}"
+proof -
+  have "\<exists>c d. a \<le> c \<and> c \<le> m \<and> m \<le> d \<and> d \<le> b \<and>
+              (\<forall>x \<in> {c..d}. f x = f m) \<and>
+              (\<forall>x \<in> {a..<c}. (f x \<noteq> f m)) \<and>
+              (\<forall>x \<in> {d<..b}. (f x \<noteq> f m)) \<and>
+              (\<forall>x \<in> {a..<c}. \<forall>y \<in> {d<..b}. f x \<noteq> f y)"
+    if m: "m \<in> {a..b}" and ab01: "{a..b} \<subseteq> {0..1}" for a b m
+  proof -
+    have comp: "compact (f -` {f m} \<inter> {0..1})"
+      by (simp add: compact_eq_bounded_closed bounded_Int closed_vimage_Int cont_f)
+    obtain c0 d0 where cd0: "{0..1} \<inter> f -` {f m} = {c0..d0}"
+      using connected_compact_interval_1 [of "{0..1} \<inter> f -` {f m}"] conn comp
+      by (metis Int_commute)
+    with that have "m \<in> cbox c0 d0"
+      by auto
+    obtain c d where cd: "{a..b} \<inter> f -` {f m} = {c..d}"
+      apply (rule_tac c="max a c0" and d="min b d0" in that)
+      using ab01 cd0 by auto
+    then have cdab: "{c..d} \<subseteq> {a..b}"
+      by blast
+    show ?thesis
+    proof (intro exI conjI ballI)
+      show "a \<le> c" "d \<le> b"
+        using cdab cd m by auto
+      show "c \<le> m" "m \<le> d"
+        using cd m by auto
+      show "\<And>x. x \<in> {c..d} \<Longrightarrow> f x = f m"
+        using cd by blast
+      show "f x \<noteq> f m" if "x \<in> {a..<c}" for x
+        using that m cd [THEN equalityD1, THEN subsetD] \<open>c \<le> m\<close> by force
+      show "f x \<noteq> f m" if "x \<in> {d<..b}" for x
+        using that m cd [THEN equalityD1, THEN subsetD, of x] \<open>m \<le> d\<close> by force
+      show "f x \<noteq> f y" if "x \<in> {a..<c}" "y \<in> {d<..b}" for x y
+      proof (cases "f x = f m \<or> f y = f m")
+        case True
+        then show ?thesis
+          using \<open>\<And>x. x \<in> {a..<c} \<Longrightarrow> f x \<noteq> f m\<close> that by auto
+      next
+        case False
+        have False if "f x = f y"
+        proof -
+          have "x \<le> m" "m \<le> y"
+            using \<open>c \<le> m\<close> \<open>x \<in> {a..<c}\<close>  \<open>m \<le> d\<close> \<open>y \<in> {d<..b}\<close> by auto
+          then have "x \<in> ({0..1} \<inter> f -` {f y})" "y \<in> ({0..1} \<inter> f -` {f y})"
+            using \<open>x \<in> {a..<c}\<close> \<open>y \<in> {d<..b}\<close> ab01 by (auto simp: that)
+          then have "m \<in> ({0..1} \<inter> f -` {f y})"
+            by (meson \<open>m \<le> y\<close> \<open>x \<le> m\<close> is_interval_connected_1 conn [of "f y"] is_interval_1)
+          with False show False by auto
+        qed
+        then show ?thesis by auto
+      qed
+    qed
+  qed
+  then obtain leftcut rightcut where LR:
+    "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow>
+            (a \<le> leftcut a b m \<and> leftcut a b m \<le> m \<and> m \<le> rightcut a b m \<and> rightcut a b m \<le> b \<and>
+            (\<forall>x \<in> {leftcut a b m..rightcut a b m}. f x = f m) \<and>
+            (\<forall>x \<in> {a..<leftcut a b m}. f x \<noteq> f m) \<and>
+            (\<forall>x \<in> {rightcut a b m<..b}. f x \<noteq> f m) \<and>
+            (\<forall>x \<in> {a..<leftcut a b m}. \<forall>y \<in> {rightcut a b m<..b}. f x \<noteq> f y))"
+    apply atomize
+    apply (clarsimp simp only: imp_conjL [symmetric] choice_iff choice_iff')
+    apply (rule that, blast)
+    done
+  then have left_right: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> a \<le> leftcut a b m \<and> rightcut a b m \<le> b"
+    and left_right_m: "\<And>a b m. \<lbrakk>m \<in> {a..b}; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> leftcut a b m \<le> m \<and> m \<le> rightcut a b m"
+    by auto
+  have left_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
+    and right_neq: "\<lbrakk>rightcut a b m < x; x \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
+    and left_right_neq: "\<lbrakk>a \<le> x; x < leftcut a b m; rightcut a b m < y; y \<le> b; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk> \<Longrightarrow> f x \<noteq> f m"
+    and feqm: "\<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk>
+                         \<Longrightarrow> f x = f m" for a b m x y
+    by (meson atLeastAtMost_iff greaterThanAtMost_iff atLeastLessThan_iff LR)+
+  have f_eqI: "\<And>a b m x y. \<lbrakk>leftcut a b m \<le> x; x \<le> rightcut a b m; leftcut a b m \<le> y; y \<le> rightcut a b m;
+                             a \<le> m; m \<le> b; {a..b} \<subseteq> {0..1}\<rbrakk>  \<Longrightarrow> f x = f y"
+    by (metis feqm)
+  define u where "u \<equiv> rightcut 0 1 0"
+  have lc[simp]: "leftcut 0 1 0 = 0" and u01: "0 \<le> u" "u \<le> 1"
+    using LR [of 0 0 1] by (auto simp: u_def)
+  have f0u: "\<And>x. x \<in> {0..u} \<Longrightarrow> f x = f 0"
+    using LR [of 0 0 1] unfolding u_def [symmetric]
+    by (metis \<open>leftcut 0 1 0 = 0\<close> atLeastAtMost_iff order_refl zero_le_one)
+  have fu1: "\<And>x. x \<in> {u<..1} \<Longrightarrow> f x \<noteq> f 0"
+    using LR [of 0 0 1] unfolding u_def [symmetric] by fastforce
+  define v where "v \<equiv> leftcut u 1 1"
+  have rc[simp]: "rightcut u 1 1 = 1" and v01: "u \<le> v" "v \<le> 1"
+    using LR [of 1 u 1] u01  by (auto simp: v_def)
+  have fuv: "\<And>x. x \<in> {u..<v} \<Longrightarrow> f x \<noteq> f 1"
+    using LR [of 1 u 1] u01 v_def by fastforce
+  have f0v: "\<And>x. x \<in> {0..<v} \<Longrightarrow> f x \<noteq> f 1"
+    by (metis f_1not0 atLeastAtMost_iff atLeastLessThan_iff f0u fuv linear)
+  have fv1: "\<And>x. x \<in> {v..1} \<Longrightarrow> f x = f 1"
+    using LR [of 1 u 1] u01 v_def by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl rc)
+  define a where "a \<equiv> leftrec u v leftcut rightcut"
+  define b where "b \<equiv> rightrec u v leftcut rightcut"
+  define c where "c \<equiv> \<lambda>x. midpoint (a x) (b x)"
+  have a_real [simp]: "a (real j) = u" for j
+    using a_def leftrec_base
+    by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral)
+  have b_real [simp]: "b (real j) = v" for j
+    using b_def rightrec_base
+    by (metis nonzero_mult_div_cancel_right of_nat_mult of_nat_numeral zero_neq_numeral)
+  have a41: "a ((4 * real m + 1) / 2^Suc n) = a ((2 * real m + 1) / 2^n)" if "n > 0" for m n
+    using that a_def leftrec_41 by blast
+  have b41: "b ((4 * real m + 1) / 2^Suc n) =
+               leftcut (a ((2 * real m + 1) / 2^n))
+                       (b ((2 * real m + 1) / 2^n))
+                       (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n
+    using that a_def b_def c_def rightrec_41 by blast
+  have a43: "a ((4 * real m + 3) / 2^Suc n) =
+               rightcut (a ((2 * real m + 1) / 2^n))
+                        (b ((2 * real m + 1) / 2^n))
+                        (c ((2 * real m + 1) / 2^n))" if "n > 0" for m n
+    using that a_def b_def c_def leftrec_43 by blast
+  have b43: "b ((4 * real m + 3) / 2^Suc n) = b ((2 * real m + 1) / 2^n)" if "n > 0" for m n
+    using that b_def rightrec_43 by blast
+  have uabv: "u \<le> a (real m / 2 ^ n) \<and> a (real m / 2 ^ n) \<le> b (real m / 2 ^ n) \<and> b (real m / 2 ^ n) \<le> v" for m n
+  proof (induction n arbitrary: m)
+    case 0
+    then show ?case by (simp add: v01)
+  next
+    case (Suc n p)
+    show ?case
+    proof (cases "even p")
+      case True
+      then obtain m where "p = 2*m" by (metis evenE)
+      then show ?thesis
+        by (simp add: Suc.IH)
+    next
+      case False
+      then obtain m where m: "p = 2*m + 1" by (metis oddE)
+      show ?thesis
+      proof (cases n)
+        case 0
+        then show ?thesis
+          by (simp add: a_def b_def leftrec_base rightrec_base v01)
+      next
+        case (Suc n')
+        then have "n > 0" by simp
+        have a_le_c: "a (real m / 2^n) \<le> c (real m / 2^n)" for m
+          unfolding c_def by (metis Suc.IH ge_midpoint_1)
+        have c_le_b: "c (real m / 2^n) \<le> b (real m / 2^n)" for m
+          unfolding c_def by (metis Suc.IH le_midpoint_1)
+        have c_ge_u: "c (real m / 2^n) \<ge> u" for m
+          using Suc.IH a_le_c order_trans by blast
+        have c_le_v: "c (real m / 2^n) \<le> v" for m
+          using Suc.IH c_le_b order_trans by blast
+        have a_ge_0: "0 \<le> a (real m / 2^n)" for m
+          using Suc.IH order_trans u01(1) by blast
+        have b_le_1: "b (real m / 2^n) \<le> 1" for m
+          using Suc.IH order_trans v01(2) by blast
+        have left_le: "leftcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<le> c ((real m) / 2^n)" for m
+          by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b)
+        have right_ge: "rightcut (a ((real m) / 2^n)) (b ((real m) / 2^n)) (c ((real m) / 2^n)) \<ge> c ((real m) / 2^n)" for m
+          by (simp add: LR a_ge_0 a_le_c b_le_1 c_le_b)
+        show ?thesis
+        proof (cases "even m")
+          case True
+          then obtain r where r: "m = 2*r" by (metis evenE)
+          show ?thesis
+            using a_le_c [of "m+1"] c_le_b [of "m+1"] a_ge_0 [of "m+1"] b_le_1 [of "m+1"]
+              Suc.IH [of "m+1"]
+            apply (simp add: r m add.commute [of 1] \<open>n > 0\<close> a41 b41 del: power_Suc)
+            apply (auto simp: left_right [THEN conjunct1])
+            using  order_trans [OF left_le c_le_v]
+            by (metis (no_types, hide_lams) add.commute mult_2 of_nat_Suc of_nat_add)
+        next
+          case False
+          then obtain r where r: "m = 2*r + 1" by (metis oddE)
+          show ?thesis
+            using a_le_c [of "m"] c_le_b [of "m"] a_ge_0 [of "m"] b_le_1 [of "m"]
+              Suc.IH [of "m+1"]
+            apply (simp add: r m add.commute [of 3] \<open>n > 0\<close> a43 b43 del: power_Suc)
+            apply (auto simp: add.commute left_right [THEN conjunct2])
+            using  order_trans [OF c_ge_u right_ge]
+             apply (metis (no_types, hide_lams) mult_2 numeral_One of_nat_add of_nat_numeral)
+            apply (metis Suc.IH mult_2 of_nat_1 of_nat_add)
+            done
+        qed
+      qed
+    qed
+  qed
+  have a_ge_0 [simp]: "0 \<le> a(m / 2^n)" and b_le_1 [simp]: "b(m / 2^n) \<le> 1" for m::nat and n
+    using uabv order_trans u01 v01 by blast+
+  then have b_ge_0 [simp]: "0 \<le> b(m / 2^n)" and a_le_1 [simp]: "a(m / 2^n) \<le> 1" for m::nat and n
+    using uabv order_trans by blast+
+  have alec [simp]: "a(m / 2^n) \<le> c(m / 2^n)" and cleb [simp]: "c(m / 2^n) \<le> b(m / 2^n)" for m::nat and n
+    by (auto simp: c_def ge_midpoint_1 le_midpoint_1 uabv)
+  have c_ge_0 [simp]: "0 \<le> c(m / 2^n)" and c_le_1 [simp]: "c(m / 2^n) \<le> 1" for m::nat and n
+    using a_ge_0 alec order_trans apply blast
+    by (meson b_le_1 cleb order_trans)
+  have "\<lbrakk>d = m-n; odd j; \<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n\<rbrakk>
+        \<Longrightarrow> (a(j / 2^n)) \<le> (c(i / 2^m)) \<and> (c(i / 2^m)) \<le> (b(j / 2^n))" for d i j m n
+  proof (induction d arbitrary: j n rule: less_induct)
+    case (less d j n)
+    show ?case
+    proof (cases "m \<le> n")
+      case True
+      have "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> = 0"
+      proof (rule Ints_nonzero_abs_less1)
+        have "(real i * 2^n - real j * 2^m) / 2^m = (real i * 2^n) / 2^m - (real j * 2^m) / 2^m"
+          using diff_divide_distrib by blast
+        also have "... = (real i * 2 ^ (n-m)) - (real j)"
+          using True by (auto simp: power_diff field_simps)
+        also have "... \<in> \<int>"
+          by simp
+        finally have "(real i * 2^n - real j * 2^m) / 2^m \<in> \<int>" .
+        with True Ints_abs show "\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar> \<in> \<int>"
+          by (fastforce simp: divide_simps)
+        show "\<bar>\<bar>2^n\<bar> * \<bar>real i / 2^m - real j / 2^n\<bar>\<bar> < 1"
+          using less.prems by (auto simp: divide_simps)
+      qed
+      then have "real i / 2^m = real j / 2^n"
+        by auto
+      then show ?thesis
+        by auto
+    next
+      case False
+      then have "n < m" by auto
+      obtain k where k: "j = Suc (2*k)"
+        using \<open>odd j\<close> oddE by fastforce
+      show ?thesis
+      proof (cases "n > 0")
+        case False
+        then have "a (real j / 2^n) = u"
+          by simp
+        also have "... \<le> c (real i / 2^m)"
+          using alec uabv by (blast intro: order_trans)
+        finally have ac: "a (real j / 2^n) \<le> c (real i / 2^m)" .
+        have "c (real i / 2^m) \<le> v"
+          using cleb uabv by (blast intro: order_trans)
+        also have "... = b (real j / 2^n)"
+          using False by simp
+        finally show ?thesis
+          by (auto simp: ac)
+      next
+        case True show ?thesis
+        proof (cases "real i / 2^m" "real j / 2^n" rule: linorder_cases)
+          case less
+          moreover have "real (4 * k + 1) / 2 ^ Suc n + 1 / (2 ^ Suc n) = real j / 2 ^ n"
+            using k by (force simp: divide_simps)
+          moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 / (2 ^ Suc n)"
+            using less.prems by simp
+          ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 1) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)"
+            using less.prems by linarith
+          have *: "a (real (4 * k + 1) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>
+                   c (real i / 2 ^ m) \<le> b (real (4 * k + 1) / 2 ^ Suc n)"
+            apply (rule less.IH [OF _ refl])
+            using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: divide_simps  \<open>n < m\<close> diff_less_mono2)
+            done
+          show ?thesis
+            using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
+            using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"]
+            using k a41 b41 * \<open>0 < n\<close>
+            apply (simp add: add.commute)
+            done
+        next
+          case equal then show ?thesis by simp
+        next
+          case greater
+          moreover have "real (4 * k + 3) / 2 ^ Suc n - 1 / (2 ^ Suc n) = real j / 2 ^ n"
+            using k by (force simp: divide_simps)
+          moreover have "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 2 * 1 / (2 ^ Suc n)"
+            using less.prems by simp
+          ultimately have closer: "\<bar>real i / 2 ^ m - real (4 * k + 3) / 2 ^ Suc n\<bar> < 1 / (2 ^ Suc n)"
+            using less.prems by linarith
+          have *: "a (real (4 * k + 3) / 2 ^ Suc n) \<le> c (real i / 2 ^ m) \<and>
+                   c (real i / 2 ^ m) \<le> b (real (4 * k + 3) / 2 ^ Suc n)"
+            apply (rule less.IH [OF _ refl])
+            using closer \<open>n < m\<close> \<open>d = m - n\<close> apply (auto simp: divide_simps  \<open>n < m\<close> diff_less_mono2)
+            done
+          show ?thesis
+            using LR [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
+            using alec [of "2*k+1"] cleb [of "2*k+1"] a_ge_0 [of "2*k+1"] b_le_1 [of "2*k+1"]
+            using k a43 b43 * \<open>0 < n\<close>
+            apply (simp add: add.commute)
+            done
+        qed
+      qed
+    qed
+  qed
+  then have aj_le_ci: "a (real j / 2 ^ n) \<le> c (real i / 2 ^ m)"
+    and ci_le_bj: "c (real i / 2 ^ m) \<le> b (real j / 2 ^ n)" if "odd j" "\<bar>real i / 2^m - real j / 2^n\<bar> < 1/2 ^ n" for i j m n
+    using that by blast+
+  have close_ab: "odd m \<Longrightarrow> \<bar>a (real m / 2 ^ n) - b (real m / 2 ^ n)\<bar> \<le> 2 / 2^n" for m n
+  proof (induction n arbitrary: m)
+    case 0
+    with u01 v01 show ?case by auto
+  next
+    case (Suc n m)
+    with oddE obtain k where k: "m = Suc (2*k)" by fastforce
+    show ?case
+    proof (cases "n > 0")
+      case False
+      with u01 v01 show ?thesis
+        by (simp add: a_def b_def leftrec_base rightrec_base)
+    next
+      case True
+      show ?thesis
+      proof (cases "even k")
+        case True
+        then obtain j where j: "k = 2*j" by (metis evenE)
+        have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n"
+        proof -
+          have "odd (Suc k)"
+            using True by auto
+          then show ?thesis
+            by (metis (no_types) Groups.add_ac(2) Suc.IH j of_nat_Suc of_nat_mult of_nat_numeral)
+        qed
+        moreover have "a ((2 * real j + 1) / 2 ^ n) \<le>
+                       leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))"
+          using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"]  b_le_1 [of "2*j+1"]
+          by (auto simp: add.commute left_right)
+        moreover have "leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le>
+                         c ((2 * real j + 1) / 2 ^ n)"
+          using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"]  b_le_1 [of "2*j+1"]
+          by (auto simp: add.commute left_right_m)
+        ultimately have "\<bar>a ((2 * real j + 1) / 2 ^ n) -
+                          leftcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))\<bar>
+                   \<le> 2/2 ^ Suc n"
+          by (simp add: c_def midpoint_def)
+        with j k \<open>n > 0\<close> show ?thesis
+          by (simp add: add.commute [of 1] a41 b41 del: power_Suc)
+      next
+        case False
+        then obtain j where j: "k = 2*j + 1" by (metis oddE)
+        have "\<bar>a ((2 * real j + 1) / 2 ^ n) - (b ((2 * real j + 1) / 2 ^ n))\<bar> \<le> 2/2 ^ n"
+          using Suc.IH [OF False] j by (auto simp: algebra_simps)
+        moreover have "c ((2 * real j + 1) / 2 ^ n) \<le>
+                       rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n))"
+          using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"]  b_le_1 [of "2*j+1"]
+          by (auto simp: add.commute left_right_m)
+        moreover have "rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) \<le>
+                         b ((2 * real j + 1) / 2 ^ n)"
+          using alec [of "2*j+1"] cleb [of "2*j+1"] a_ge_0 [of "2*j+1"]  b_le_1 [of "2*j+1"]
+          by (auto simp: add.commute left_right)
+        ultimately have "\<bar>rightcut (a ((2 * real j + 1) / 2 ^ n)) (b ((2 * real j + 1) / 2 ^ n)) (c ((2 * real j + 1) / 2 ^ n)) -
+                          b ((2 * real j + 1) / 2 ^ n)\<bar>  \<le>  2/2 ^ Suc n"
+          by (simp add: c_def midpoint_def)
+        with j k \<open>n > 0\<close> show ?thesis
+          by (simp add: add.commute [of 3] a43 b43 del: power_Suc)
+      qed
+    qed
+  qed
+  have m1_to_3: "4 * real k - 1 = real (4 * (k-1)) + 3" if "0 < k" for k
+    using that by auto
+  have fb_eq_fa: "\<lbrakk>0 < j; 2*j < 2 ^ n\<rbrakk> \<Longrightarrow> f(b((2 * real j - 1) / 2^n)) = f(a((2 * real j + 1) / 2^n))" for n j
+  proof (induction n arbitrary: j)
+    case 0
+    then show ?case by auto
+  next
+    case (Suc n j) show ?case
+    proof (cases "n > 0")
+      case False
+      with Suc.prems show ?thesis by auto
+    next
+      case True
+      show ?thesis proof (cases "even j")
+        case True
+        then obtain k where k: "j = 2*k" by (metis evenE)
+        with \<open>0 < j\<close> have "k > 0" "2 * k < 2 ^ n"
+          using Suc.prems(2) k by auto
+        with k \<open>0 < n\<close> Suc.IH [of k] show ?thesis
+          apply (simp add: m1_to_3 a41 b43 del: power_Suc)
+          apply (subst of_nat_diff, auto)
+          done
+      next
+        case False
+        then obtain k where k: "j = 2*k + 1" by (metis oddE)
+        have "f (leftcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))
+              = f (c ((2 * k + 1) / 2^n))"
+          "f (c ((2 * k + 1) / 2^n))
+              = f (rightcut (a ((2 * k + 1) / 2^n)) (b ((2 * k + 1) / 2^n)) (c ((2 * k + 1) / 2^n)))"
+          using alec [of "2*k+1" n] cleb [of "2*k+1" n] a_ge_0 [of "2*k+1" n]  b_le_1 [of "2*k+1" n] k
+          using left_right_m [of "c((2*k + 1) / 2^n)" "a((2*k + 1) / 2^n)" "b((2*k + 1) / 2^n)"]
+           apply (auto simp: add.commute  feqm [OF order_refl]  feqm [OF _ order_refl, symmetric])
+          done
+        then
+        show ?thesis
+          by (simp add: k add.commute [of 1] add.commute [of 3] a43 b41\<open>0 < n\<close> del: power_Suc)
+      qed
+    qed
+  qed
+  have f_eq_fc: "\<lbrakk>0 < j; j < 2 ^ n\<rbrakk>
+                 \<Longrightarrow> f(b((2*j - 1) / 2 ^ (Suc n))) = f(c(j / 2^n)) \<and>
+                     f(a((2*j + 1) / 2 ^ (Suc n))) = f(c(j / 2^n))" for n and j::nat
+  proof (induction n arbitrary: j)
+    case 0
+    then show ?case by auto
+  next
+    case (Suc n)
+    show ?case
+    proof (cases "even j")
+      case True
+      then obtain k where k: "j = 2*k" by (metis evenE)
+      then have less2n: "k < 2 ^ n"
+        using Suc.prems(2) by auto
+      have "0 < k" using \<open>0 < j\<close> k by linarith
+      then have m1_to_3: "real (4 * k - Suc 0) = real (4 * (k-1)) + 3"
+        by auto
+      then show ?thesis
+        using Suc.IH [of k] k \<open>0 < k\<close>
+        apply (simp add: less2n add.commute [of 1] m1_to_3 a41 b43 del: power_Suc)
+        apply (auto simp: of_nat_diff)
+        done
+    next
+      case False
+      then obtain k where k: "j = 2*k + 1" by (metis oddE)
+      with Suc.prems have "k < 2^n" by auto
+      show ?thesis
+        using alec [of "2*k+1" "Suc n"] cleb [of "2*k+1" "Suc n"] a_ge_0 [of "2*k+1" "Suc n"]  b_le_1 [of "2*k+1" "Suc n"] k
+        using left_right_m [of "c((2*k + 1) / 2 ^ Suc n)" "a((2*k + 1) / 2 ^ Suc n)" "b((2*k + 1) / 2 ^ Suc n)"]
+        apply (simp add: add.commute [of 1] add.commute [of 3] m1_to_3 b41 a43 del: power_Suc)
+        apply (force intro: feqm)
+        done
+    qed
+  qed
+  define D01 where "D01 \<equiv> {0<..<1} \<inter> (\<Union>k m. {real m / 2^k})"
+  have cloD01 [simp]: "closure D01 = {0..1}"
+    unfolding D01_def
+    by (subst closure_dyadic_rationals_in_convex_set_pos_1) auto
+  have "uniformly_continuous_on D01 (f \<circ> c)"
+  proof (clarsimp simp: uniformly_continuous_on_def)
+    fix e::real
+    assume "0 < e"
+    have ucontf: "uniformly_continuous_on {0..1} f"
+      by (simp add: compact_uniformly_continuous [OF cont_f])
+    then obtain d where "0 < d" and d: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; norm (x' - x) < d\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e/2"
+      unfolding uniformly_continuous_on_def dist_norm
+      by (metis \<open>0 < e\<close> less_divide_eq_numeral1(1) mult_zero_left)
+    obtain n where n: "1/2^n < min d 1"
+      by (metis \<open>0 < d\<close> divide_less_eq_1 less_numeral_extra(1) min_def one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_numeral)
+    with gr0I have "n > 0"
+      by (force simp: divide_simps)
+    show "\<exists>d>0. \<forall>x\<in>D01. \<forall>x'\<in>D01. dist x' x < d \<longrightarrow> dist (f (c x')) (f (c x)) < e"
+    proof (intro exI ballI impI conjI)
+      show "(0::real) < 1/2^n" by auto
+    next
+      have dist_fc_close: "dist (f(c(real i / 2^m))) (f(c(real j / 2^n))) < e/2"
+        if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and clo: "abs(i / 2^m - j / 2^n) < 1/2 ^ n" for i j m
+      proof -
+        have abs3: "\<bar>x - a\<bar> < e \<Longrightarrow> x = a \<or> \<bar>x - (a - e/2)\<bar> < e/2 \<or> \<bar>x - (a + e/2)\<bar> < e/2" for x a e::real
+          by linarith
+        consider "i / 2 ^ m = j / 2 ^ n"
+          | "\<bar>i / 2 ^ m - (2 * j - 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n"
+          | "\<bar>i / 2 ^ m - (2 * j + 1) / 2 ^ Suc n\<bar> < 1/2 ^ Suc n"
+          using abs3 [OF clo] j by (auto simp: field_simps of_nat_diff)
+        then show ?thesis
+        proof cases
+          case 1 with \<open>0 < e\<close> show ?thesis by auto
+        next
+          case 2
+          have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> b - c < d" for a b c
+            by auto
+          have "norm (c (real i / 2 ^ m) - b (real (2 * j - 1) / 2 ^ Suc n)) < d"
+            using 2 j n close_ab [of "2*j-1" "Suc n"]
+            using b_ge_0 [of "2*j-1" "Suc n"] b_le_1 [of "2*j-1" "Suc n"]
+            using aj_le_ci [of "2*j-1" i m "Suc n"]
+            using ci_le_bj [of "2*j-1" i m "Suc n"]
+            apply (simp add: divide_simps of_nat_diff del: power_Suc)
+            apply (auto simp: divide_simps intro!: *)
+            done
+          moreover have "f(c(j / 2^n)) = f(b ((2*j - 1) / 2 ^ (Suc n)))"
+            using f_eq_fc [OF j] by metis
+          ultimately show ?thesis
+            by (metis dist_norm atLeastAtMost_iff b_ge_0 b_le_1 c_ge_0 c_le_1 d)
+        next
+          case 3
+          have *: "abs(a - b) \<le> 1/2 ^ n \<and> 1/2 ^ n < d \<and> a \<le> c \<and> c \<le> b \<Longrightarrow> c - a < d" for a b c
+            by auto
+          have "norm (c (real i / 2 ^ m) - a (real (2 * j + 1) / 2 ^ Suc n)) < d"
+            using 3 j n close_ab [of "2*j+1" "Suc n"]
+            using b_ge_0 [of "2*j+1" "Suc n"] b_le_1 [of "2*j+1" "Suc n"]
+            using aj_le_ci [of "2*j+1" i m "Suc n"]
+            using ci_le_bj [of "2*j+1" i m "Suc n"]
+            apply (simp add: divide_simps of_nat_diff del: power_Suc)
+            apply (auto simp: divide_simps intro!: *)
+            done
+          moreover have "f(c(j / 2^n)) = f(a ((2*j + 1) / 2 ^ (Suc n)))"
+            using f_eq_fc [OF j] by metis
+          ultimately show ?thesis
+            by (metis dist_norm a_ge_0 atLeastAtMost_iff a_ge_0 a_le_1 c_ge_0 c_le_1 d)
+        qed
+      qed
+      show "dist (f (c x')) (f (c x)) < e"
+        if "x \<in> D01" "x' \<in> D01" "dist x' x < 1/2^n" for x x'
+        using that unfolding D01_def dyadics_in_open_unit_interval
+      proof clarsimp
+        fix i k::nat and m p
+        assume i: "0 < i" "i < 2 ^ m" and k: "0<k" "k < 2 ^ p"
+        assume clo: "dist (real k / 2 ^ p) (real i / 2 ^ m) < 1/2 ^ n"
+        obtain j::nat where "0 < j" "j < 2 ^ n"
+          and clo_ij: "abs(i / 2^m - j / 2^n) < 1/2 ^ n"
+          and clo_kj: "abs(k / 2^p - j / 2^n) < 1/2 ^ n"
+        proof -
+          have "max (2^n * i / 2^m) (2^n * k / 2^p) \<ge> 0"
+            by (auto simp: le_max_iff_disj)
+          then obtain j where "floor (max (2^n*i / 2^m) (2^n*k / 2^p)) = int j"
+            using zero_le_floor zero_le_imp_eq_int by blast
+          then have j_le: "real j \<le> max (2^n * i / 2^m) (2^n * k / 2^p)"
+            and less_j1: "max (2^n * i / 2^m) (2^n * k / 2^p) < real j + 1"
+            using floor_correct [of "max (2^n * i / 2^m) (2^n * k / 2^p)"] by linarith+
+          show thesis
+          proof (cases "j = 0")
+            case True
+            show thesis
+            proof
+              show "(1::nat) < 2 ^ n"
+                apply (subst one_less_power)
+                using \<open>n > 0\<close> by auto
+              show "\<bar>real i / 2 ^ m - real 1/2 ^ n\<bar> < 1/2 ^ n"
+                using i less_j1 by (simp add: dist_norm field_simps True)
+              show "\<bar>real k / 2 ^ p - real 1/2 ^ n\<bar> < 1/2 ^ n"
+                using k less_j1 by (simp add: dist_norm field_simps True)
+            qed simp
+          next
+            case False
+            have 1: "real j * 2 ^ m < real i * 2 ^ n"
+              if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * 2 ^ m < real i * 2 ^ p"
+              for i k m p
+            proof -
+              have "real j * 2 ^ p * 2 ^ m \<le> real k * 2 ^ n * 2 ^ m"
+                using j by simp
+              moreover have "real k * 2 ^ m * 2 ^ n < real i * 2 ^ p * 2 ^ n"
+                using k by simp
+              ultimately have "real j * 2 ^ p * 2 ^ m < real i * 2 ^ p * 2 ^ n"
+                by (simp only: mult_ac)
+              then show ?thesis
+                by simp
+            qed
+            have 2: "real j * 2 ^ m < 2 ^ m + real i * 2 ^ n"
+              if j: "real j * 2 ^ p \<le> real k * 2 ^ n" and k: "real k * (2 ^ m * 2 ^ n) < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)"
+              for i k m p
+            proof -
+              have "real j * 2 ^ p * 2 ^ m \<le> real k * (2 ^ m * 2 ^ n)"
+                using j by simp
+              also have "... < 2 ^ m * 2 ^ p + real i * (2 ^ n * 2 ^ p)"
+                by (rule k)
+              finally have "(real j * 2 ^ m) * 2 ^ p < (2 ^ m + real i * 2 ^ n) * 2 ^ p"
+                by (simp add: algebra_simps)
+              then show ?thesis
+                by simp
+            qed
+            have 3: "real j * 2 ^ p < 2 ^ p + real k * 2 ^ n"
+              if j: "real j * 2 ^ m \<le> real i * 2 ^ n" and i: "real i * 2 ^ p \<le> real k * 2 ^ m"
+            proof -
+              have "real j * 2 ^ m * 2 ^ p \<le> real i * 2 ^ n * 2 ^ p"
+                using j by simp
+              moreover have "real i * 2 ^ p * 2 ^ n \<le> real k * 2 ^ m * 2 ^ n"
+                using i by simp
+              ultimately have "real j * 2 ^ m * 2 ^ p \<le> real k * 2 ^ m * 2 ^ n"
+                by (simp only: mult_ac)
+              then have "real j * 2 ^ p \<le> real k * 2 ^ n"
+                by simp
+              also have "... < 2 ^ p + real k * 2 ^ n"
+                by auto
+              finally show ?thesis by simp
+            qed
+            show ?thesis
+            proof
+              have "real j < 2 ^ n"
+                using j_le i k
+                apply (auto simp: le_max_iff_disj simp del: real_of_nat_less_numeral_power_cancel_iff elim!: le_less_trans)
+                 apply (auto simp: field_simps)
+                done
+              then show "j < 2 ^ n"
+                by auto
+              show "\<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2 ^ n"
+                using clo less_j1 j_le
+                apply (auto simp: le_max_iff_disj divide_simps dist_norm)
+                 apply (auto simp: algebra_simps abs_if split: if_split_asm dest: 1 2)
+                done
+              show "\<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2 ^ n"
+                using  clo less_j1 j_le
+                apply (auto simp: le_max_iff_disj divide_simps dist_norm)
+                 apply (auto simp: algebra_simps not_less abs_if split: if_split_asm dest: 3 2)
+                done
+            qed (use False in simp)
+          qed
+        qed
+        show "dist (f (c (real k / 2 ^ p))) (f (c (real i / 2 ^ m))) < e"
+        proof (rule dist_triangle_half_l)
+          show "dist (f (c (real k / 2 ^ p))) (f(c(j / 2^n))) < e/2"
+            apply (rule dist_fc_close)
+            using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> k clo_kj by auto
+          show "dist (f (c (real i / 2 ^ m))) (f (c (real j / 2 ^ n))) < e/2"
+            apply (rule dist_fc_close)
+            using \<open>0 < j\<close> \<open>j < 2 ^ n\<close> i clo_ij by auto
+        qed
+      qed
+    qed
+  qed
+  then obtain h where ucont_h: "uniformly_continuous_on {0..1} h"
+    and fc_eq: "\<And>x. x \<in> D01 \<Longrightarrow> (f \<circ> c) x = h x"
+  proof (rule uniformly_continuous_on_extension_on_closure [of D01 "f \<circ> c"])
+  qed (use closure_subset [of D01] in \<open>auto intro!: that\<close>)
+  then have cont_h: "continuous_on {0..1} h"
+    using uniformly_continuous_imp_continuous by blast
+  have h_eq: "h (real k / 2 ^ m) = f (c (real k / 2 ^ m))" if "0 < k" "k < 2^m" for k m
+    using fc_eq that by (force simp: D01_def)
+  have "h ` {0..1} = f ` {0..1}"
+  proof
+    have "h ` (closure D01) \<subseteq> f ` {0..1}"
+    proof (rule image_closure_subset)
+      show "continuous_on (closure D01) h"
+        using cont_h by simp
+      show "closed (f ` {0..1})"
+        using compact_continuous_image [OF cont_f] compact_imp_closed by blast
+      show "h ` D01 \<subseteq> f ` {0..1}"
+        by (force simp: dyadics_in_open_unit_interval D01_def h_eq)
+    qed
+    with cloD01 show "h ` {0..1} \<subseteq> f ` {0..1}" by simp
+    have a12 [simp]: "a (1/2) = u"
+      by (metis a_def leftrec_base numeral_One of_nat_numeral)
+    have b12 [simp]: "b (1/2) = v"
+      by (metis b_def rightrec_base numeral_One of_nat_numeral)
+    have "f ` {0..1} \<subseteq> closure(h ` D01)"
+    proof (clarsimp simp: closure_approachable dyadics_in_open_unit_interval D01_def)
+      fix x e::real
+      assume "0 \<le> x" "x \<le> 1" "0 < e"
+      have ucont_f: "uniformly_continuous_on {0..1} f"
+        using compact_uniformly_continuous cont_f by blast
+      then obtain \<delta> where "\<delta> > 0"
+        and \<delta>: "\<And>x x'. \<lbrakk>x \<in> {0..1}; x' \<in> {0..1}; dist x' x < \<delta>\<rbrakk> \<Longrightarrow> norm (f x' - f x) < e"
+        using \<open>0 < e\<close> by (auto simp: uniformly_continuous_on_def dist_norm)
+      have *: "\<exists>m::nat. \<exists>y. odd m \<and> 0 < m \<and> m < 2 ^ n \<and> y \<in> {a(m / 2^n) .. b(m / 2^n)} \<and> f y = f x"
+        if "n \<noteq> 0" for n
+        using that
+      proof (induction n)
+        case 0 then show ?case by auto
+      next
+        case (Suc n)
+        show ?case
+        proof (cases "n=0")
+          case True
+          consider "x \<in> {0..u}" | "x \<in> {u..v}" | "x \<in> {v..1}"
+            using \<open>0 \<le> x\<close> \<open>x \<le> 1\<close> by force
+          then have "\<exists>y\<ge>a (real 1/2). y \<le> b (real 1/2) \<and> f y = f x"
+          proof cases
+            case 1
+            then show ?thesis
+              apply (rule_tac x=u in exI)
+              using uabv [of 1 1] f0u [of u] f0u [of x] by auto
+          next
+            case 2
+            then show ?thesis
+              by (rule_tac x=x in exI) auto
+          next
+            case 3
+            then show ?thesis
+              apply (rule_tac x=v in exI)
+              using uabv [of 1 1] fv1 [of v] fv1 [of x] by auto
+          qed
+          with \<open>n=0\<close> show ?thesis
+            by (rule_tac x=1 in exI) auto
+        next
+          case False
+          with Suc obtain m y
+            where "odd m" "0 < m" and mless: "m < 2 ^ n"
+              and y: "y \<in> {a (real m / 2 ^ n)..b (real m / 2 ^ n)}" and feq: "f y = f x"
+            by metis
+          then obtain j where j: "m = 2*j + 1" by (metis oddE)
+          consider "y \<in> {a((2*j + 1) / 2^n) .. b((4*j + 1) / 2 ^ (Suc n))}"
+            | "y \<in> {b((4*j + 1) / 2 ^ (Suc n)) .. a((4*j + 3) / 2 ^ (Suc n))}"
+            | "y \<in> {a((4*j + 3) / 2 ^ (Suc n)) .. b((2*j + 1) / 2^n)}"
+            using y j by force
+          then show ?thesis
+          proof cases
+            case 1
+            then show ?thesis
+              apply (rule_tac x="4*j + 1" in exI)
+              apply (rule_tac x=y in exI)
+              using mless j \<open>n \<noteq> 0\<close>
+              apply (simp add: feq a41 b41 add.commute [of 1] del: power_Suc)
+              apply (simp add: algebra_simps)
+              done
+          next
+            case 2
+            show ?thesis
+              apply (rule_tac x="4*j + 1" in exI)
+              apply (rule_tac x="b((4*j + 1) / 2 ^ (Suc n))" in exI)
+              using mless \<open>n \<noteq> 0\<close> 2 j
+              using alec [of "2*j+1" n] cleb [of "2*j+1" n] a_ge_0 [of "2*j+1" n]  b_le_1 [of "2*j+1" n]
+              using left_right [of "c((2*j + 1) / 2^n)" "a((2*j + 1) / 2^n)" "b((2*j + 1) / 2^n)"]
+              apply (simp add: a41 b41 a43 b43 add.commute [of 1] add.commute [of 3] del: power_Suc)
+              apply (auto simp: feq [symmetric] intro: f_eqI)
+              done
+          next
+            case 3
+            then show ?thesis
+              apply (rule_tac x="4*j + 3" in exI)
+              apply (rule_tac x=y in exI)
+              using mless j \<open>n \<noteq> 0\<close>
+              apply (simp add: feq a43 b43 del: power_Suc)
+              apply (simp add: algebra_simps)
+              done
+          qed
+        qed
+      qed
+      obtain n where n: "1/2^n < min (\<delta> / 2) 1"
+        by (metis \<open>0 < \<delta>\<close> divide_less_eq_1 less_numeral_extra(1) min_less_iff_conj one_less_numeral_iff power_one_over real_arch_pow_inv semiring_norm(76) zero_less_divide_iff zero_less_numeral)
+      with gr0I have "n \<noteq> 0"
+        by fastforce
+      with * obtain m::nat and y
+        where "odd m" "0 < m" and mless: "m < 2 ^ n"
+          and y: "y \<in> {a(m / 2^n) .. b(m / 2^n)}" and feq: "f x = f y"
+        by metis
+      then have "0 \<le> y" "y \<le> 1"
+        by (metis atLeastAtMost_iff a_ge_0 b_le_1 order.trans)+
+      moreover have "y < \<delta> + c (real m / 2 ^ n)" "c (real m / 2 ^ n) < \<delta> + y"
+        using y apply simp_all
+        using alec [of m n] cleb [of m n] n real_sum_of_halves close_ab [OF \<open>odd m\<close>, of n]
+        by linarith+
+      moreover note \<open>0 < m\<close> mless \<open>0 \<le> x\<close> \<open>x \<le> 1\<close>
+      ultimately show "\<exists>k. \<exists>m\<in>{0<..<2 ^ k}. dist (h (real m / 2 ^ k)) (f x) < e"
+        apply (rule_tac x=n in exI)
+        apply (rule_tac x=m in bexI)
+         apply (auto simp: dist_norm h_eq feq \<delta>)
+        done
+    qed
+    also have "... \<subseteq> h ` {0..1}"
+      apply (rule closure_minimal)
+      using compact_continuous_image [OF cont_h] compact_imp_closed by (auto simp: D01_def)
+    finally show "f ` {0..1} \<subseteq> h ` {0..1}" .
+  qed
+  moreover have "inj_on h {0..1}"
+  proof -
+    have "u < v"
+      by (metis atLeastAtMost_iff f0u f_1not0 fv1 order.not_eq_order_implies_strict u01(1) u01(2) v01(1))
+    have f_not_fu: "\<And>x. \<lbrakk>u < x; x \<le> v\<rbrakk> \<Longrightarrow> f x \<noteq> f u"
+      by (metis atLeastAtMost_iff f0u fu1 greaterThanAtMost_iff order_refl order_trans u01(1) v01(2))
+    have f_not_fv: "\<And>x. \<lbrakk>u \<le> x; x < v\<rbrakk> \<Longrightarrow> f x \<noteq> f v"
+      by (metis atLeastAtMost_iff order_refl order_trans v01(2) atLeastLessThan_iff fuv fv1)
+    have a_less_b:
+         "a(j / 2^n) < b(j / 2^n) \<and>
+          (\<forall>x. a(j / 2^n) < x \<longrightarrow> x \<le> b(j / 2^n) \<longrightarrow> f x \<noteq> f(a(j / 2^n))) \<and>
+          (\<forall>x. a(j / 2^n) \<le> x \<longrightarrow> x < b(j / 2^n) \<longrightarrow> f x \<noteq> f(b(j / 2^n)))" for n and j::nat
+    proof (induction n arbitrary: j)
+      case 0 then show ?case
+        by (simp add: \<open>u < v\<close> f_not_fu f_not_fv)
+    next
+      case (Suc n j) show ?case
+      proof (cases "n > 0")
+        case False then show ?thesis
+          by (auto simp: a_def b_def leftrec_base rightrec_base \<open>u < v\<close> f_not_fu f_not_fv)
+      next
+        case True show ?thesis
+        proof (cases "even j")
+          case True
+          with \<open>0 < n\<close> Suc.IH show ?thesis
+            by (auto elim!: evenE)
+        next
+          case False
+          then obtain k where k: "j = 2*k + 1" by (metis oddE)
+          then show ?thesis
+          proof (cases "even k")
+            case True
+            then obtain m where m: "k = 2*m" by (metis evenE)
+            have fleft: "f (leftcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) =
+                         f (c((2*m + 1) / 2^n))"
+              using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n]  b_le_1 [of "2*m+1" n]
+              using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+              by (auto intro: f_eqI)
+            show ?thesis
+            proof (intro conjI impI notI allI)
+              have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)"
+              proof -
+                have "f (c ((1 + real m * 2) / 2 ^ n)) = f (a ((1 + real m * 2) / 2 ^ n))"
+                  using k m \<open>0 < n\<close> fleft that a41 [of n m] b41 [of n m]
+                  using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n]  b_le_1 [of "2*m+1" n]
+                  using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+                  by (auto simp: algebra_simps)
+                moreover have "a (real (1 + m * 2) / 2 ^ n) < c (real (1 + m * 2) / 2 ^ n)"
+                  using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def)
+                moreover have "c (real (1 + m * 2) / 2 ^ n) \<le> b (real (1 + m * 2) / 2 ^ n)"
+                  using cleb by blast
+                ultimately show ?thesis
+                  using Suc.IH [of "1 + m * 2"] by force
+              qed
+              then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force
+            next
+              fix x
+              assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))"
+              then show False
+                using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct1]
+                using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m]
+                using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
+                using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+                by (auto simp: algebra_simps)
+            next
+              fix x
+              assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))"
+              then show False
+                using k m \<open>0 < n\<close> a41 [of n m] b41 [of n m] fleft left_neq
+                using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
+                by (auto simp: algebra_simps)
+            qed
+          next
+            case False
+            with oddE obtain m where m: "k = Suc (2*m)" by fastforce
+            have fright: "f (rightcut (a ((2*m + 1) / 2^n)) (b ((2*m + 1) / 2^n)) (c ((2*m + 1) / 2^n))) = f (c((2*m + 1) / 2^n))"
+              using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n]  b_le_1 [of "2*m+1" n]
+              using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+              by (auto intro: f_eqI [OF _ order_refl])
+            show ?thesis
+            proof (intro conjI impI notI allI)
+              have False if "b (real j / 2 ^ Suc n) \<le> a (real j / 2 ^ Suc n)"
+              proof -
+                have "f (c ((1 + real m * 2) / 2 ^ n)) = f (b ((1 + real m * 2) / 2 ^ n))"
+                  using k m \<open>0 < n\<close> fright that a43 [of n m] b43 [of n m]
+                  using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n]  b_le_1 [of "2*m+1" n]
+                  using left_right [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+                  by (auto simp: algebra_simps)
+                moreover have "a (real (1 + m * 2) / 2 ^ n) \<le> c (real (1 + m * 2) / 2 ^ n)"
+                  using alec by blast
+                moreover have "c (real (1 + m * 2) / 2 ^ n) < b (real (1 + m * 2) / 2 ^ n)"
+                  using Suc.IH [of "1 + m * 2"] by (simp add: c_def midpoint_def)
+                ultimately show ?thesis
+                  using Suc.IH [of "1 + m * 2"] by force
+              qed
+              then show "a (real j / 2 ^ Suc n) < b (real j / 2 ^ Suc n)" by force
+            next
+              fix x
+              assume "a (real j / 2 ^ Suc n) < x" "x \<le> b (real j / 2 ^ Suc n)" "f x = f (a (real j / 2 ^ Suc n))"
+              then show False
+                using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m] fright right_neq
+                using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
+                by (auto simp: algebra_simps)
+            next
+              fix x
+              assume "a (real j / 2 ^ Suc n) \<le> x" "x < b (real j / 2 ^ Suc n)" "f x = f (b (real j / 2 ^ Suc n))"
+              then show False
+                using Suc.IH [of "1 + m * 2", THEN conjunct2, THEN conjunct2]
+                using k m \<open>0 < n\<close> a43 [of n m] b43 [of n m]
+                using alec [of "2*m+1" n] cleb [of "2*m+1" n] a_ge_0 [of "2*m+1" n] b_le_1 [of "2*m+1" n]
+                using left_right_m [of "c((2*m + 1) / 2^n)" "a((2*m + 1) / 2^n)" "b((2*m + 1) / 2^n)"]
+                by (auto simp: algebra_simps fright simp del: power_Suc)
+            qed
+          qed
+        qed
+      qed
+    qed
+    have c_gt_0 [simp]: "0 < c(m / 2^n)" and c_less_1 [simp]: "c(m / 2^n) < 1" for m::nat and n
+      using a_less_b [of m n] apply (simp_all add: c_def midpoint_def)
+      using a_ge_0 [of m n] b_le_1 [of m n] apply linarith+
+      done
+    have approx: "\<exists>j n. odd j \<and> n \<noteq> 0 \<and>
+                        real i / 2^m \<le> real j / 2^n \<and>
+                        real j / 2^n \<le> real k / 2^p \<and>
+                        \<bar>real i / 2 ^ m - real j / 2 ^ n\<bar> < 1/2^n \<and>
+                        \<bar>real k / 2 ^ p - real j / 2 ^ n\<bar> < 1/2^n"
+      if "0 < i" "i < 2 ^ m" "0 < k" "k < 2 ^ p" "i / 2^m < k / 2^p" "m + p = N" for N m p i k
+      using that
+    proof (induction N arbitrary: m p i k rule: less_induct)
+      case (less N)
+      then consider "i / 2^m \<le> 1/2" "1/2 \<le> k / 2^p" | "k / 2^p < 1/2" | "k / 2^p \<ge> 1/2" "1/2 < i / 2^m"
+        by linarith
+      then show ?case
+      proof cases
+        case 1
+        with less.prems show ?thesis
+          by (rule_tac x=1 in exI)+ (fastforce simp: divide_simps)
+      next
+        case 2 show ?thesis
+        proof (cases m)
+          case 0 with less.prems show ?thesis
+            by auto
+        next
+          case (Suc m') show ?thesis
+          proof (cases p)
+            case 0 with less.prems show ?thesis by auto
+          next
+            case (Suc p')
+            have False if "real i * 2 ^ p' < real k * 2 ^ m'" "k < 2 ^ p'" "2 ^ m' \<le> i"
+            proof -
+              have "real k * 2 ^ m' < 2 ^ p' * 2 ^ m'"
+                using that by simp
+              then have "real i * 2 ^ p' < 2 ^ p' * 2 ^ m'"
+                using that by linarith
+              with that show ?thesis by simp
+            qed
+            then show ?thesis
+              using less.IH [of "m'+p'" i m' k p'] less.prems \<open>m = Suc m'\<close> 2 Suc
+              apply atomize
+              apply (force simp: divide_simps)
+              done
+          qed
+        qed
+      next
+        case 3 show ?thesis
+        proof (cases m)
+          case 0 with less.prems show ?thesis
+            by auto
+        next
+          case (Suc m') show ?thesis
+          proof (cases p)
+            case 0 with less.prems show ?thesis by auto
+          next
+            case (Suc p')
+            then show ?thesis
+              using less.IH [of "m'+p'" "i - 2^m'" m' "k - 2 ^ p'" p'] less.prems \<open>m = Suc m'\<close> Suc 3
+              apply atomize
+              apply (auto simp: field_simps of_nat_diff)
+              apply (rule_tac x="2 ^ n + j" in exI, simp)
+              apply (rule_tac x="Suc n" in exI)
+              apply (auto simp: field_simps)
+              done
+          qed
+        qed
+      qed
+    qed
+    have clec: "c(real i / 2^m) \<le> c(real j / 2^n)"
+      if i: "0 < i" "i < 2 ^ m" and j: "0 < j" "j < 2 ^ n" and ij: "i / 2^m < j / 2^n" for m i n j
+    proof -
+      obtain j' n' where "odd j'" "n' \<noteq> 0"
+        and i_le_j: "real i / 2 ^ m \<le> real j' / 2 ^ n'"
+        and j_le_j: "real j' / 2 ^ n' \<le> real j / 2 ^ n"
+        and clo_ij: "\<bar>real i / 2 ^ m - real j' / 2 ^ n'\<bar> < 1/2 ^ n'"
+        and clo_jj: "\<bar>real j / 2 ^ n - real j' / 2 ^ n'\<bar> < 1/2 ^ n'"
+        using approx [of i m j n "m+n"] that i j ij by auto
+      with oddE obtain q where q: "j' = Suc (2*q)" by fastforce
+      have "c (real i / 2 ^ m) \<le> c((2*q + 1) / 2^n')"
+      proof (cases "i / 2^m = (2*q + 1) / 2^n'")
+        case True then show ?thesis by simp
+      next
+        case False
+        with i_le_j q have less: "i / 2^m < (2*q + 1) / 2^n'"
+          by auto
+        have *: "\<lbrakk>i < q; abs(i - q) < s*2; q = r + s\<rbrakk> \<Longrightarrow> abs(i - r) < s" for i q s r::real
+          by auto
+        have "c(i / 2^m) \<le> b(real(4 * q + 1) / 2 ^ (Suc n'))"
+          apply (rule ci_le_bj, force)
+          apply (rule * [OF less])
+          using i_le_j clo_ij q apply (auto simp: divide_simps)
+          done
+        then show ?thesis
+          using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] b41 [of n' q] \<open>n' \<noteq> 0\<close>
+          using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"]
+          by (auto simp: algebra_simps)
+      qed
+      also have "... \<le> c(real j / 2^n)"
+      proof (cases "j / 2^n = (2*q + 1) / 2^n'")
+        case True
+        then show ?thesis by simp
+      next
+        case False
+        with j_le_j q have less: "(2*q + 1) / 2^n' < j / 2^n"
+          by auto
+        have *: "\<lbrakk>q < i; abs(i - q) < s*2; r = q + s\<rbrakk> \<Longrightarrow> abs(i - r) < s" for i q s r::real
+          by auto
+        have "a(real(4*q + 3) / 2 ^ (Suc n')) \<le> c(j / 2^n)"
+          apply (rule aj_le_ci, force)
+          apply (rule * [OF less])
+          using j_le_j clo_jj q apply (auto simp: divide_simps)
+          done
+        then show ?thesis
+          using alec [of "2*q+1" n'] cleb [of "2*q+1" n'] a_ge_0 [of "2*q+1" n'] b_le_1 [of "2*q+1" n'] a43 [of n' q] \<open>n' \<noteq> 0\<close>
+          using left_right_m [of "c((2*q + 1) / 2^n')" "a((2*q + 1) / 2^n')" "b((2*q + 1) / 2^n')"]
+          by (auto simp: algebra_simps)
+      qed
+      finally show ?thesis .
+    qed
+    have "x = y" if "0 \<le> x" "x \<le> 1" "0 \<le> y" "y \<le> 1" "h x = h y" for x y
+      using that
+    proof (induction x y rule: linorder_class.linorder_less_wlog)
+      case (less x1 x2)
+      obtain m n where m: "0 < m" "m < 2 ^ n"
+        and x12: "x1 < m / 2^n" "m / 2^n < x2"
+        and neq: "h x1 \<noteq> h (real m / 2^n)"
+      proof -
+        have "(x1 + x2) / 2 \<in> closure D01"
+          using cloD01 less.hyps less.prems by auto
+        with less obtain y where "y \<in> D01" and dist_y: "dist y ((x1 + x2) / 2) < (x2 - x1) / 64"
+          unfolding closure_approachable
+          by (metis diff_gt_0_iff_gt less_divide_eq_numeral1(1) mult_zero_left)
+        obtain m n where m: "0 < m" "m < 2 ^ n"
+                     and clo: "\<bar>real m / 2 ^ n - (x1 + x2) / 2\<bar> < (x2 - x1) / 64"
+                     and n: "1/2^n < (x2 - x1) / 128"
+        proof -
+          have "min 1 ((x2 - x1) / 128) > 0" "1/2 < (1::real)"
+            using less by auto
+          then obtain N where N: "1/2^N < min 1 ((x2 - x1) / 128)"
+            by (metis power_one_over real_arch_pow_inv)
+          then have "N > 0"
+            using less_divide_eq_1 by force
+          obtain p q where p: "p < 2 ^ q" "p \<noteq> 0" and yeq: "y = real p / 2 ^ q"
+            using \<open>y \<in> D01\<close> by (auto simp: zero_less_divide_iff D01_def)
+          show ?thesis
+          proof
+            show "0 < 2^N * p"
+              using p by auto
+            show "2 ^ N * p < 2 ^ (N+q)"
+              by (simp add: p power_add)
+            have "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> = \<bar>real p / 2 ^ q - (x1 + x2) / 2\<bar>"
+              by (simp add: power_add)
+            also have "... = \<bar>y - (x1 + x2) / 2\<bar>"
+              by (simp add: yeq)
+            also have "... < (x2 - x1) / 64"
+              using dist_y by (simp add: dist_norm)
+            finally show "\<bar>real (2 ^ N * p) / 2 ^ (N + q) - (x1 + x2) / 2\<bar> < (x2 - x1) / 64" .
+            have "(1::real) / 2 ^ (N + q) \<le> 1/2^N"
+              by (simp add: field_simps)
+            also have "... < (x2 - x1) / 128"
+              using N by force
+            finally show "1/2 ^ (N + q) < (x2 - x1) / 128" .
+          qed
+        qed
+        obtain m' n' m'' n'' where "0 < m'" "m' < 2 ^ n'" "x1 < m' / 2^n'" "m' / 2^n' < x2"
+          and "0 < m''" "m'' < 2 ^ n''" "x1 < m'' / 2^n''" "m'' / 2^n'' < x2"
+          and neq: "h (real m'' / 2^n'') \<noteq> h (real m' / 2^n')"
+        proof
+          show "0 < Suc (2*m)"
+            by simp
+          show m21: "Suc (2*m) < 2 ^ Suc n"
+            using m by auto
+          show "x1 < real (Suc (2 * m)) / 2 ^ Suc n"
+            using clo by (simp add: field_simps abs_if split: if_split_asm)
+          show "real (Suc (2 * m)) / 2 ^ Suc n < x2"
+            using n clo by (simp add: field_simps abs_if split: if_split_asm)
+          show "0 < 4*m + 3"
+            by simp
+          have "m+1 \<le> 2 ^ n"
+            using m by simp
+          then have "4 * (m+1) \<le> 4 * (2 ^ n)"
+            by simp
+          then show m43: "4*m + 3 < 2 ^ (n+2)"
+            by (simp add: algebra_simps)
+          show "x1 < real (4 * m + 3) / 2 ^ (n + 2)"
+            using clo by (simp add: field_simps abs_if split: if_split_asm)
+          show "real (4 * m + 3) / 2 ^ (n + 2) < x2"
+            using n clo by (simp add: field_simps abs_if split: if_split_asm)
+          have c_fold: "midpoint (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n)) = c ((2 * real m + 1) / 2 ^ Suc n)"
+            by (simp add: c_def)
+          define R where "R \<equiv> rightcut (a ((2 * real m + 1) / 2 ^ Suc n)) (b ((2 * real m + 1) / 2 ^ Suc n))  (c ((2 * real m + 1) / 2 ^ Suc n))"
+          have "R < b ((2 * real m + 1) / 2 ^ Suc n)"
+            unfolding R_def  using a_less_b [of "4*m + 3" "n+2"] a43 [of "Suc n" m] b43 [of "Suc n" m]
+            by simp
+          then have Rless: "R < midpoint R (b ((2 * real m + 1) / 2 ^ Suc n))"
+            by (simp add: midpoint_def)
+          have midR_le: "midpoint R (b ((2 * real m + 1) / 2 ^ Suc n)) \<le> b ((2 * real m + 1) / (2 * 2 ^ n))"
+            using \<open>R < b ((2 * real m + 1) / 2 ^ Suc n)\<close>
+            by (simp add: midpoint_def)
+          have "(real (Suc (2 * m)) / 2 ^ Suc n) \<in> D01"  "real (4 * m + 3) / 2 ^ (n + 2) \<in> D01"
+            by (simp_all add: D01_def m21 m43 del: power_Suc of_nat_Suc of_nat_add add_2_eq_Suc') blast+
+          then show "h (real (4 * m + 3) / 2 ^ (n + 2)) \<noteq> h (real (Suc (2 * m)) / 2 ^ Suc n)"
+            using a_less_b [of "4*m + 3" "n+2", THEN conjunct1]
+            using a43 [of "Suc n" m] b43 [of "Suc n" m]
+            using alec [of "2*m+1" "Suc n"] cleb [of "2*m+1" "Suc n"] a_ge_0 [of "2*m+1" "Suc n"]  b_le_1 [of "2*m+1" "Suc n"]
+            apply (simp add: fc_eq [symmetric] c_def del: power_Suc)
+            apply (simp only: add.commute [of 1] c_fold R_def [symmetric])
+            apply (rule right_neq)
+            using Rless apply (simp add: R_def)
+               apply (rule midR_le, auto)
+            done
+        qed
+        then show ?thesis by (metis that)
+      qed
+      have m_div: "0 < m / 2^n" "m / 2^n < 1"
+        using m  by (auto simp: divide_simps)
+      have closure0m: "{0..m / 2^n} = closure ({0<..< m / 2^n} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
+        by (subst closure_dyadic_rationals_in_convex_set_pos_1, simp_all add: not_le m)
+      have closurem1: "{m / 2^n .. 1} = closure ({m / 2^n <..< 1} \<inter> (\<Union>k m. {real m / 2 ^ k}))"
+        apply (subst closure_dyadic_rationals_in_convex_set_pos_1; simp add: not_le m)
+        using \<open>0 < real m / 2 ^ n\<close> by linarith
+      have cont_h': "continuous_on (closure ({u<..<v} \<inter> (\<Union>k m. {real m / 2 ^ k}))) h"
+        if "0 \<le> u" "v \<le> 1" for u v
+        apply (rule continuous_on_subset [OF cont_h])
+        apply (rule closure_minimal [OF subsetI])
+        using that apply auto
+        done
+      have closed_f': "closed (f ` {u..v})" if "0 \<le> u" "v \<le> 1" for u v
+        by (metis compact_continuous_image cont_f compact_interval atLeastatMost_subset_iff
+            compact_imp_closed continuous_on_subset that)
+      have less_2I: "\<And>k i. real i / 2 ^ k < 1 \<Longrightarrow> i < 2 ^ k"
+        by simp
+      have "h ` ({0<..<m / 2 ^ n} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {0..c (m / 2 ^ n)}"
+      proof clarsimp
+        fix p q
+        assume p: "0 < real p / 2 ^ q" "real p / 2 ^ q < real m / 2 ^ n"
+        then have [simp]: "0 < p" "p < 2 ^ q"
+           apply (simp add: divide_simps)
+          apply (blast intro: p less_2I m_div less_trans)
+          done
+        have "f (c (real p / 2 ^ q)) \<in> f ` {0..c (real m / 2 ^ n)}"
+          by (auto simp: clec p m)
+        then show "h (real p / 2 ^ q) \<in> f ` {0..c (real m / 2 ^ n)}"
+          by (simp add: h_eq)
+      qed
+      then have "h ` {0 .. m / 2^n} \<subseteq> f ` {0 .. c(m / 2^n)}"
+        apply (subst closure0m)
+        apply (rule image_closure_subset [OF cont_h' closed_f'])
+        using m_div apply auto
+        done
+      then have hx1: "h x1 \<in> f ` {0 .. c(m / 2^n)}"
+        using x12 less.prems(1) by auto
+      then obtain t1 where t1: "h x1 = f t1" "0 \<le> t1" "t1 \<le> c (m / 2 ^ n)"
+        by auto
+      have "h ` ({m / 2 ^ n<..<1} \<inter> (\<Union>q p. {real p / 2 ^ q})) \<subseteq> f ` {c (m / 2 ^ n)..1}"
+      proof clarsimp
+        fix p q
+        assume p: "real m / 2 ^ n < real p / 2 ^ q" and [simp]: "p < 2 ^ q"
+        then have [simp]: "0 < p"
+          using gr_zeroI m_div by fastforce
+        have "f (c (real p / 2 ^ q)) \<in> f ` {c (m / 2 ^ n)..1}"
+          by (auto simp: clec p m)
+        then show "h (real p / 2 ^ q) \<in> f ` {c (real m / 2 ^ n)..1}"
+          by (simp add: h_eq)
+      qed
+      then have "h ` {m / 2^n .. 1} \<subseteq> f ` {c(m / 2^n) .. 1}"
+        apply (subst closurem1)
+        apply (rule image_closure_subset [OF cont_h' closed_f'])
+        using m apply auto
+        done
+      then have hx2: "h x2 \<in> f ` {c(m / 2^n)..1}"
+        using x12 less.prems by auto
+      then obtain t2 where t2: "h x2 = f t2" "c (m / 2 ^ n) \<le> t2" "t2 \<le> 1"
+        by auto
+      with t1 less neq have False
+        using conn [of "h x2", unfolded is_interval_connected_1 [symmetric] is_interval_1, rule_format, of t1 t2 "c(m / 2^n)"]
+        by (simp add: h_eq m)
+      then show ?case by blast
+    qed auto
+    then show ?thesis
+      by (auto simp: inj_on_def)
+  qed
+  ultimately have "{0..1::real} homeomorphic f ` {0..1}"
+    using homeomorphic_compact [OF _ cont_h] by blast
+  then show ?thesis
+    using homeomorphic_sym by blast
+qed
+
+
+theorem path_contains_arc:
+  fixes p :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
+  assumes "path p" and a: "pathstart p = a" and b: "pathfinish p = b" and "a \<noteq> b"
+  obtains q where "arc q" "path_image q \<subseteq> path_image p" "pathstart q = a" "pathfinish q = b"
+proof -
+  have ucont_p: "uniformly_continuous_on {0..1} p"
+    using \<open>path p\<close> unfolding path_def
+    by (metis compact_Icc compact_uniformly_continuous)
+  define \<phi> where "\<phi> \<equiv> \<lambda>S. S \<subseteq> {0..1} \<and> 0 \<in> S \<and> 1 \<in> S \<and>
+                           (\<forall>x \<in> S. \<forall>y \<in> S. open_segment x y \<inter> S = {} \<longrightarrow> p x = p y)"
+  obtain T where "closed T" "\<phi> T" and T: "\<And>U. \<lbrakk>closed U; \<phi> U\<rbrakk> \<Longrightarrow> \<not> (U \<subset> T)"
+  proof (rule Brouwer_reduction_theorem_gen [of "{0..1}" \<phi>])
+    have *: "{x<..<y} \<inter> {0..1} = {x<..<y}" if "0 \<le> x" "y \<le> 1" "x \<le> y" for x y::real
+      using that by auto
+    show "\<phi> {0..1}"
+      by (auto simp: \<phi>_def open_segment_eq_real_ivl *)
+    show "\<phi> (INTER UNIV F)"
+      if "\<And>n. closed (F n)" and \<phi>: "\<And>n. \<phi> (F n)" and Fsub: "\<And>n. F (Suc n) \<subseteq> F n" for F
+    proof -
+      have F01: "\<And>n. F n \<subseteq> {0..1} \<and> 0 \<in> F n \<and> 1 \<in> F n"
+        and peq: "\<And>n x y. \<lbrakk>x \<in> F n; y \<in> F n; open_segment x y \<inter> F n = {}\<rbrakk> \<Longrightarrow> p x = p y"
+        by (metis \<phi> \<phi>_def)+
+      have pqF: False if "\<forall>u. x \<in> F u" "\<forall>x. y \<in> F x" "open_segment x y \<inter> (\<Inter>x. F x) = {}" and neg: "p x \<noteq> p y"
+        for x y
+        using that
+      proof (induction x y rule: linorder_class.linorder_less_wlog)
+        case (less x y)
+        have xy: "x \<in> {0..1}" "y \<in> {0..1}"
+          by (metis less.prems subsetCE F01)+
+        have "norm(p x - p y) / 2 > 0"
+          using less by auto
+        then obtain e where "e > 0"
+          and e: "\<And>u v. \<lbrakk>u \<in> {0..1}; v \<in> {0..1}; dist v u < e\<rbrakk> \<Longrightarrow> dist (p v) (p u) < norm(p x - p y) / 2"
+          by (metis uniformly_continuous_onE [OF ucont_p])
+        have minxy: "min e (y - x)  < (y - x) * (3 / 2)"
+          by (subst min_less_iff_disj) (simp add: less)
+        obtain w z where "w < z" and w: "w \<in> {x<..<y}" and z: "z \<in> {x<..<y}"
+          and wxe: "norm(w - x) < e" and zye: "norm(z - y) < e"
+          apply (rule_tac w = "x + (min e (y - x) / 3)" and z = "y - (min e (y - x) / 3)" in that)
+          using minxy \<open>0 < e\<close> less by simp_all
+        have Fclo: "\<And>T. T \<in> range F \<Longrightarrow> closed T"
+          by (metis \<open>\<And>n. closed (F n)\<close> image_iff)
+        have eq: "{w..z} \<inter> INTER UNIV F = {}"
+          using less w z apply (auto simp: open_segment_eq_real_ivl)
+          by (metis (no_types, hide_lams) INT_I IntI empty_iff greaterThanLessThan_iff not_le order.trans)
+        then obtain K where "finite K" and K: "{w..z} \<inter> (\<Inter> (F ` K)) = {}"
+          by (metis finite_subset_image compact_imp_fip [OF compact_interval Fclo])
+        then have "K \<noteq> {}"
+          using \<open>w < z\<close> \<open>{w..z} \<inter> INTER K F = {}\<close> by auto
+        define n where "n \<equiv> Max K"
+        have "n \<in> K" unfolding n_def by (metis \<open>K \<noteq> {}\<close> \<open>finite K\<close> Max_in)
+        have "F n \<subseteq> \<Inter> (F ` K)"
+          unfolding n_def by (metis Fsub Max_ge \<open>K \<noteq> {}\<close> \<open>finite K\<close> cINF_greatest lift_Suc_antimono_le)
+        with K have wzF_null: "{w..z} \<inter> F n = {}"
+          by (metis disjoint_iff_not_equal subset_eq)
+        obtain u where u: "u \<in> F n" "u \<in> {x..w}" "({u..w} - {u}) \<inter> F n = {}"
+        proof (cases "w \<in> F n")
+          case True
+          then show ?thesis
+            by (metis wzF_null \<open>w < z\<close> atLeastAtMost_iff disjoint_iff_not_equal less_eq_real_def)
+        next
+          case False
+          obtain u where "u \<in> F n" "u \<in> {x..w}" "{u<..<w} \<inter> F n = {}"
+          proof (rule segment_to_point_exists [of "F n \<inter> {x..w}" w])
+            show "closed (F n \<inter> {x..w})"
+              by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_real_atLeastAtMost)
+            show "F n \<inter> {x..w} \<noteq> {}"
+              by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(1) less_eq_real_def w)
+          qed (auto simp: open_segment_eq_real_ivl intro!: that)
+          with False show thesis
+            apply (auto simp: disjoint_iff_not_equal intro!: that)
+            by (metis greaterThanLessThan_iff less_eq_real_def)
+        qed
+        obtain v where v: "v \<in> F n" "v \<in> {z..y}" "({z..v} - {v}) \<inter> F n = {}"
+        proof (cases "z \<in> F n")
+          case True
+          have "z \<in> {w..z}"
+            using \<open>w < z\<close> by auto
+          then show ?thesis
+            by (metis wzF_null Int_iff True empty_iff)
+        next
+          case False
+          show ?thesis
+          proof (rule segment_to_point_exists [of "F n \<inter> {z..y}" z])
+            show "closed (F n \<inter> {z..y})"
+              by (metis \<open>\<And>n. closed (F n)\<close> closed_Int closed_atLeastAtMost)
+            show "F n \<inter> {z..y} \<noteq> {}"
+              by (metis atLeastAtMost_iff disjoint_iff_not_equal greaterThanLessThan_iff less.prems(2) less_eq_real_def z)
+            show "\<And>b. \<lbrakk>b \<in> F n \<inter> {z..y}; open_segment z b \<inter> (F n \<inter> {z..y}) = {}\<rbrakk> \<Longrightarrow> thesis"
+              apply (rule that)
+                apply (auto simp: open_segment_eq_real_ivl)
+              by (metis DiffI Int_iff atLeastAtMost_diff_ends atLeastAtMost_iff atLeastatMost_empty_iff empty_iff insert_iff False)
+          qed
+        qed
+        obtain u v where "u \<in> {0..1}" "v \<in> {0..1}" "norm(u - x) < e" "norm(v - y) < e" "p u = p v"
+        proof
+          show "u \<in> {0..1}" "v \<in> {0..1}"
+            by (metis F01 \<open>u \<in> F n\<close> \<open>v \<in> F n\<close> subsetD)+
+          show "norm(u - x) < e" "norm (v - y) < e"
+            using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> atLeastAtMost_iff real_norm_def wxe zye by auto
+          show "p u = p v"
+          proof (rule peq)
+            show "u \<in> F n" "v \<in> F n"
+              by (auto simp: u v)
+            have "False" if "\<xi> \<in> F n" "u < \<xi>" "\<xi> < v" for \<xi>
+            proof -
+              have "\<xi> \<notin> {z..v}"
+                by (metis DiffI disjoint_iff_not_equal less_irrefl singletonD that v(3))
+              moreover have "\<xi> \<notin> {w..z} \<inter> F n"
+                by (metis equals0D wzF_null)
+              ultimately have "\<xi> \<in> {u..w}"
+                using that by auto
+              then show ?thesis
+                by (metis DiffI disjoint_iff_not_equal less_eq_real_def not_le singletonD that u(3))
+            qed
+            moreover
+            have "\<lbrakk>\<xi> \<in> F n; v < \<xi>; \<xi> < u\<rbrakk> \<Longrightarrow> False" for \<xi>
+              using \<open>u \<in> {x..w}\<close> \<open>v \<in> {z..y}\<close> \<open>w < z\<close> by simp
+            ultimately
+            show "open_segment u v \<inter> F n = {}"
+              by (force simp: open_segment_eq_real_ivl)
+          qed
+        qed
+        then show ?case
+          using e [of x u] e [of y v] xy
+          apply (simp add: open_segment_eq_real_ivl dist_norm del: divide_const_simps)
+          by (metis dist_norm dist_triangle_half_r less_irrefl)
+      qed (auto simp: open_segment_commute)
+      show ?thesis
+        unfolding \<phi>_def by (metis (no_types, hide_lams) INT_I Inf_lower2 rangeI that F01 subsetCE pqF)
+    qed
+    show "closed {0..1::real}" by auto
+  qed (meson \<phi>_def)
+  then have "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T"
+    and peq: "\<And>x y. \<lbrakk>x \<in> T; y \<in> T; open_segment x y \<inter> T = {}\<rbrakk> \<Longrightarrow> p x = p y"
+    unfolding \<phi>_def by metis+
+  then have "T \<noteq> {}" by auto
+  define h where "h \<equiv> \<lambda>x. p(@y. y \<in> T \<and> open_segment x y \<inter> T = {})"
+  have "p y = p z" if "y \<in> T" "z \<in> T" and xyT: "open_segment x y \<inter> T = {}" and xzT: "open_segment x z \<inter> T = {}"
+    for x y z
+  proof (cases "x \<in> T")
+    case True
+    with that show ?thesis by (metis \<open>\<phi> T\<close> \<phi>_def)
+  next
+    case False
+    have "insert x (open_segment x y \<union> open_segment x z) \<inter> T = {}"
+      by (metis False Int_Un_distrib2 Int_insert_left Un_empty_right xyT xzT)
+    moreover have "open_segment y z \<inter> T \<subseteq> insert x (open_segment x y \<union> open_segment x z) \<inter> T"
+      apply auto
+      by (metis greaterThanLessThan_iff less_eq_real_def less_le_trans linorder_neqE_linordered_idom open_segment_eq_real_ivl)
+    ultimately have "open_segment y z \<inter> T = {}"
+      by blast
+    with that peq show ?thesis by metis
+  qed
+  then have h_eq_p_gen: "h x = p y" if "y \<in> T" "open_segment x y \<inter> T = {}" for x y
+    using that unfolding h_def
+    by (metis (mono_tags, lifting) some_eq_ex)
+  then have h_eq_p: "\<And>x. x \<in> T \<Longrightarrow> h x = p x"
+    by simp
+  have disjoint: "\<And>x. \<exists>y. y \<in> T \<and> open_segment x y \<inter> T = {}"
+    by (meson \<open>T \<noteq> {}\<close> \<open>closed T\<close> segment_to_point_exists)
+  have heq: "h x = h x'" if "open_segment x x' \<inter> T = {}" for x x'
+  proof (cases "x \<in> T \<or> x' \<in> T")
+    case True
+    then show ?thesis
+      by (metis h_eq_p h_eq_p_gen open_segment_commute that)
+  next
+    case False
+    obtain y y' where "y \<in> T" "open_segment x y \<inter> T = {}" "h x = p y"
+      "y' \<in> T" "open_segment x' y' \<inter> T = {}" "h x' = p y'"
+      by (meson disjoint h_eq_p_gen)
+    moreover have "open_segment y y' \<subseteq> (insert x (insert x' (open_segment x y \<union> open_segment x' y' \<union> open_segment x x')))"
+      by (auto simp: open_segment_eq_real_ivl)
+    ultimately show ?thesis
+      using False that by (fastforce simp add: h_eq_p intro!: peq)
+  qed
+  have "h ` {0..1} homeomorphic {0..1::real}"
+  proof (rule homeomorphic_monotone_image_interval)
+    show "continuous_on {0..1} h"
+    proof (clarsimp simp add: continuous_on_iff)
+      fix u \<epsilon>::real
+      assume "0 < \<epsilon>" "0 \<le> u" "u \<le> 1"
+      then obtain \<delta> where "\<delta> > 0" and \<delta>: "\<And>v. v \<in> {0..1} \<Longrightarrow> dist v u < \<delta> \<longrightarrow> dist (p v) (p u) < \<epsilon> / 2"
+        using ucont_p [unfolded uniformly_continuous_on_def]
+        by (metis atLeastAtMost_iff half_gt_zero_iff)
+      then have "dist (h v) (h u) < \<epsilon>" if "v \<in> {0..1}" "dist v u < \<delta>" for v
+      proof (cases "open_segment u v \<inter> T = {}")
+        case True
+        then show ?thesis
+          using \<open>0 < \<epsilon>\<close> heq by auto
+      next
+        case False
+        have uvT: "closed (closed_segment u v \<inter> T)" "closed_segment u v \<inter> T \<noteq> {}"
+          using False open_closed_segment by (auto simp: \<open>closed T\<close> closed_Int)
+        obtain w where "w \<in> T" and w: "w \<in> closed_segment u v" "open_segment u w \<inter> T = {}"
+          apply (rule segment_to_point_exists [OF uvT, of u])
+          by (metis IntD1 Int_commute Int_left_commute ends_in_segment(1) inf.orderE subset_oc_segment)
+        then have puw: "dist (p u) (p w) < \<epsilon> / 2"
+          by (metis (no_types) \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
+        obtain z where "z \<in> T" and z: "z \<in> closed_segment u v" "open_segment v z \<inter> T = {}"
+          apply (rule segment_to_point_exists [OF uvT, of v])
+          by (metis IntD2 Int_commute Int_left_commute ends_in_segment(2) inf.orderE subset_oc_segment)
+        then have "dist (p u) (p z) < \<epsilon> / 2"
+          by (metis \<open>T \<subseteq> {0..1}\<close> \<open>dist v u < \<delta>\<close> \<delta> dist_commute dist_in_closed_segment le_less_trans subsetCE)
+        then show ?thesis
+          using puw by (metis (no_types) \<open>w \<in> T\<close> \<open>z \<in> T\<close> dist_commute dist_triangle_half_l h_eq_p_gen w(2) z(2))
+      qed
+      with \<open>0 < \<delta>\<close> show "\<exists>\<delta>>0. \<forall>v\<in>{0..1}. dist v u < \<delta> \<longrightarrow> dist (h v) (h u) < \<epsilon>" by blast
+    qed
+    show "connected ({0..1} \<inter> h -` {z})" for z
+    proof (clarsimp simp add: connected_iff_connected_component)
+      fix u v
+      assume huv_eq: "h v = h u" and uv: "0 \<le> u" "u \<le> 1" "0 \<le> v" "v \<le> 1"
+      have "\<exists>T. connected T \<and> T \<subseteq> {0..1} \<and> T \<subseteq> h -` {h u} \<and> u \<in> T \<and> v \<in> T"
+      proof (intro exI conjI)
+        show "connected (closed_segment u v)"
+          by simp
+        show "closed_segment u v \<subseteq> {0..1}"
+          by (simp add: uv closed_segment_eq_real_ivl)
+        have pxy: "p x = p y"
+          if "T \<subseteq> {0..1}" "0 \<in> T" "1 \<in> T" "x \<in> T" "y \<in> T"
+          and disjT: "open_segment x y \<inter> (T - open_segment u v) = {}"
+          and xynot: "x \<notin> open_segment u v" "y \<notin> open_segment u v"
+          for x y
+        proof (cases "open_segment x y \<inter> open_segment u v = {}")
+          case True
+          then show ?thesis
+            by (metis Diff_Int_distrib Diff_empty peq disjT \<open>x \<in> T\<close> \<open>y \<in> T\<close>)
+        next
+          case False
+          then have "open_segment x u \<union> open_segment y v \<subseteq> open_segment x y - open_segment u v \<or>
+                     open_segment y u \<union> open_segment x v \<subseteq> open_segment x y - open_segment u v" (is "?xuyv \<or> ?yuxv")
+            using xynot by (fastforce simp add: open_segment_eq_real_ivl not_le not_less split: if_split_asm)
+          then show "p x = p y"
+          proof
+            assume "?xuyv"
+            then have "open_segment x u \<inter> T = {}" "open_segment y v \<inter> T = {}"
+              using disjT by auto
+            then have "h x = h y"
+              using heq huv_eq by auto
+            then show ?thesis
+              using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto
+          next
+            assume "?yuxv"
+            then have "open_segment y u \<inter> T = {}" "open_segment x v \<inter> T = {}"
+              using disjT by auto
+            then have "h x = h y"
+              using heq [of y u]  heq [of x v] huv_eq by auto
+            then show ?thesis
+              using h_eq_p \<open>x \<in> T\<close> \<open>y \<in> T\<close> by auto
+          qed
+        qed
+        have "\<not> T - open_segment u v \<subset> T"
+        proof (rule T)
+          show "closed (T - open_segment u v)"
+            by (simp add: closed_Diff [OF \<open>closed T\<close>] open_segment_eq_real_ivl)
+          have "0 \<notin> open_segment u v" "1 \<notin> open_segment u v"
+            using open_segment_eq_real_ivl uv by auto
+          then show "\<phi> (T - open_segment u v)"
+            using \<open>T \<subseteq> {0..1}\<close> \<open>0 \<in> T\<close> \<open>1 \<in> T\<close>
+            by (auto simp: \<phi>_def) (meson peq pxy)
+        qed
+        then have "open_segment u v \<inter> T = {}"
+          by blast
+        then show "closed_segment u v \<subseteq> h -` {h u}"
+          by (force intro: heq simp: open_segment_eq_real_ivl closed_segment_eq_real_ivl split: if_split_asm)+
+      qed auto
+      then show "connected_component ({0..1} \<inter> h -` {h u}) u v"
+        by (simp add: connected_component_def)
+    qed
+    show "h 1 \<noteq> h 0"
+      by (metis \<open>\<phi> T\<close> \<phi>_def a \<open>a \<noteq> b\<close> b h_eq_p pathfinish_def pathstart_def)
+  qed
+  then obtain f and g :: "real \<Rightarrow> 'a"
+    where gfeq: "(\<forall>x\<in>h ` {0..1}. (g(f x) = x))" and fhim: "f ` h ` {0..1} = {0..1}" and contf: "continuous_on (h ` {0..1}) f"
+      and fgeq: "(\<forall>y\<in>{0..1}. (f(g y) = y))" and pag: "path_image g = h ` {0..1}" and contg: "continuous_on {0..1} g"
+    by (auto simp: homeomorphic_def homeomorphism_def path_image_def)
+  then have "arc g"
+    by (metis arc_def path_def inj_on_def)
+  obtain u v where "u \<in> {0..1}" "a = g u" "v \<in> {0..1}" "b = g v"
+    by (metis (mono_tags, hide_lams) \<open>\<phi> T\<close> \<phi>_def a b fhim gfeq h_eq_p imageI path_image_def pathfinish_def pathfinish_in_path_image pathstart_def pathstart_in_path_image)
+  then have "a \<in> path_image g" "b \<in> path_image g"
+    using path_image_def by blast+
+  have ph: "path_image h \<subseteq> path_image p"
+    by (metis image_mono image_subset_iff path_image_def disjoint h_eq_p_gen \<open>T \<subseteq> {0..1}\<close>)
+  show ?thesis
+  proof
+    show "pathstart (subpath u v g) = a" "pathfinish (subpath u v g) = b"
+      by (simp_all add: \<open>a = g u\<close> \<open>b = g v\<close>)
+    show "path_image (subpath u v g) \<subseteq> path_image p"
+      by (metis \<open>arc g\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_imp_path order_trans pag path_image_def path_image_subpath_subset ph)
+    show "arc (subpath u v g)"
+      using \<open>arc g\<close> \<open>a = g u\<close> \<open>b = g v\<close> \<open>u \<in> {0..1}\<close> \<open>v \<in> {0..1}\<close> arc_subpath_arc \<open>a \<noteq> b\<close> by blast
+  qed
+qed
+
+
+corollary path_connected_arcwise:
+  fixes S :: "'a::{complete_space,real_normed_vector} set"
+  shows "path_connected S \<longleftrightarrow>
+         (\<forall>x \<in> S. \<forall>y \<in> S. x \<noteq> y \<longrightarrow> (\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y))"
+        (is "?lhs = ?rhs")
+proof (intro iffI impI ballI)
+  fix x y
+  assume "path_connected S" "x \<in> S" "y \<in> S" "x \<noteq> y"
+  then obtain p where p: "path p" "path_image p \<subseteq> S" "pathstart p = x" "pathfinish p = y"
+    by (force simp: path_connected_def)
+  then show "\<exists>g. arc g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
+    by (metis \<open>x \<noteq> y\<close> order_trans path_contains_arc)
+next
+  assume R [rule_format]: ?rhs
+  show ?lhs
+    unfolding path_connected_def
+  proof (intro ballI)
+    fix x y
+    assume "x \<in> S" "y \<in> S"
+    show "\<exists>g. path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
+    proof (cases "x = y")
+      case True with \<open>x \<in> S\<close> path_component_def path_component_refl show ?thesis
+        by blast
+    next
+      case False with R [OF \<open>x \<in> S\<close> \<open>y \<in> S\<close>] show ?thesis
+        by (auto intro: arc_imp_path)
+    qed
+  qed
+qed
+
+
+corollary arc_connected_trans:
+  fixes g :: "real \<Rightarrow> 'a::{complete_space,real_normed_vector}"
+  assumes "arc g" "arc h" "pathfinish g = pathstart h" "pathstart g \<noteq> pathfinish h"
+  obtains i where "arc i" "path_image i \<subseteq> path_image g \<union> path_image h"
+                  "pathstart i = pathstart g" "pathfinish i = pathfinish h"
+  by (metis (no_types, hide_lams) arc_imp_path assms path_contains_arc path_image_join path_join pathfinish_join pathstart_join)
+
+
+
+
+subsection\<open>Accessibility of frontier points\<close>
+
+lemma dense_accessible_frontier_points:
+  fixes S :: "'a::{complete_space,real_normed_vector} set"
+  assumes "open S" and opeSV: "openin (subtopology euclidean (frontier S)) V" and "V \<noteq> {}"
+  obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
+proof -
+  obtain z where "z \<in> V"
+    using \<open>V \<noteq> {}\<close> by auto
+  then obtain r where "r > 0" and r: "ball z r \<inter> frontier S \<subseteq> V"
+    by (metis openin_contains_ball opeSV)
+  then have "z \<in> frontier S"
+    using \<open>z \<in> V\<close> opeSV openin_contains_ball by blast
+  then have "z \<in> closure S" "z \<notin> S"
+    by (simp_all add: frontier_def assms interior_open)
+  with \<open>r > 0\<close> have "infinite (S \<inter> ball z r)"
+    by (auto simp: closure_def islimpt_eq_infinite_ball)
+  then obtain y where "y \<in> S" and y: "y \<in> ball z r"
+    using infinite_imp_nonempty by force
+  then have "y \<notin> frontier S"
+    by (meson \<open>open S\<close> disjoint_iff_not_equal frontier_disjoint_eq)
+  have "y \<noteq> z"
+    using \<open>y \<in> S\<close> \<open>z \<notin> S\<close> by blast
+  have "path_connected(ball z r)"
+    by (simp add: convex_imp_path_connected)
+  with y \<open>r > 0\<close>  obtain g where "arc g" and pig: "path_image g \<subseteq> ball z r"
+                                 and g: "pathstart g = y" "pathfinish g = z"
+    using \<open>y \<noteq> z\<close> by (force simp: path_connected_arcwise)
+  have "compact (g -` frontier S \<inter> {0..1})"
+    apply (simp add: compact_eq_bounded_closed bounded_Int bounded_closed_interval)
+     apply (rule closed_vimage_Int)
+    using \<open>arc g\<close> apply (auto simp: arc_def path_def)
+    done
+  moreover have "g -` frontier S \<inter> {0..1} \<noteq> {}"
+  proof -
+    have "\<exists>r. r \<in> g -` frontier S \<and> r \<in> {0..1}"
+      by (metis \<open>z \<in> frontier S\<close> g(2) imageE path_image_def pathfinish_in_path_image vimageI2)
+    then show ?thesis
+      by blast
+  qed
+  ultimately obtain t where gt: "g t \<in> frontier S" and "0 \<le> t" "t \<le> 1"
+                and t: "\<And>u. \<lbrakk>g u \<in> frontier S; 0 \<le> u; u \<le> 1\<rbrakk> \<Longrightarrow> t \<le> u"
+    by (force simp: dest!: compact_attains_inf)
+  moreover have "t \<noteq> 0"
+    by (metis \<open>y \<notin> frontier S\<close> g(1) gt pathstart_def)
+  ultimately have  t01: "0 < t" "t \<le> 1"
+    by auto
+  have "V \<subseteq> frontier S"
+    using opeSV openin_contains_ball by blast
+  show ?thesis
+  proof
+    show "arc (subpath 0 t g)"
+      by (simp add: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> \<open>arc g\<close> \<open>t \<noteq> 0\<close> arc_subpath_arc)
+    have "g 0 \<in> S"
+      by (metis \<open>y \<in> S\<close> g(1) pathstart_def)
+    then show "pathstart (subpath 0 t g) \<in> S"
+      by auto
+    have "g t \<in> V"
+      by (metis IntI atLeastAtMost_iff gt image_eqI path_image_def pig r subsetCE \<open>0 \<le> t\<close> \<open>t \<le> 1\<close>)
+    then show "pathfinish (subpath 0 t g) \<in> V"
+      by auto
+    then have "inj_on (subpath 0 t g) {0..1}"
+      using t01
+      apply (clarsimp simp: inj_on_def subpath_def)
+      apply (drule inj_onD [OF arc_imp_inj_on [OF \<open>arc g\<close>]])
+      using mult_le_one apply auto
+      done
+    then have "subpath 0 t g ` {0..<1} \<subseteq> subpath 0 t g ` {0..1} - {subpath 0 t g 1}"
+      by (force simp: dest: inj_onD)
+    moreover have False if "subpath 0 t g ` ({0..<1}) - S \<noteq> {}"
+    proof -
+      have contg: "continuous_on {0..1} g"
+        using \<open>arc g\<close> by (auto simp: arc_def path_def)
+      have "subpath 0 t g ` {0..<1} \<inter> frontier S \<noteq> {}"
+      proof (rule connected_Int_frontier [OF _ _ that])
+        show "connected (subpath 0 t g ` {0..<1})"
+          apply (rule connected_continuous_image)
+           apply (simp add: subpath_def)
+           apply (intro continuous_intros continuous_on_compose2 [OF contg])
+           apply (auto simp: \<open>0 \<le> t\<close> \<open>t \<le> 1\<close> mult_le_one)
+          done
+        show "subpath 0 t g ` {0..<1} \<inter> S \<noteq> {}"
+          using \<open>y \<in> S\<close> g(1) by (force simp: subpath_def image_def pathstart_def)
+      qed
+      then obtain x where "x \<in> subpath 0 t g ` {0..<1}" "x \<in> frontier S"
+        by blast
+      with t01 \<open>0 \<le> t\<close> mult_le_one t show False
+        by (fastforce simp: subpath_def)
+    qed
+    then have "subpath 0 t g ` {0..1} - {subpath 0 t g 1} \<subseteq> S"
+      using subsetD by fastforce
+    ultimately  show "subpath 0 t g ` {0..<1} \<subseteq> S"
+      by auto
+  qed
+qed
+
+
+lemma dense_accessible_frontier_points_connected:
+  fixes S :: "'a::{complete_space,real_normed_vector} set"
+  assumes "open S" "connected S" "x \<in> S" "V \<noteq> {}"
+      and ope: "openin (subtopology euclidean (frontier S)) V"
+  obtains g where "arc g" "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
+proof -
+  have "V \<subseteq> frontier S"
+    using ope openin_imp_subset by blast
+  with \<open>open S\<close> \<open>x \<in> S\<close> have "x \<notin> V"
+    using interior_open by (auto simp: frontier_def)
+  obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g \<in> S" "pathfinish g \<in> V"
+    by (metis dense_accessible_frontier_points [OF \<open>open S\<close> ope \<open>V \<noteq> {}\<close>])
+  then have "path_connected S"
+    by (simp add: assms connected_open_path_connected)
+  with \<open>pathstart g \<in> S\<close> \<open>x \<in> S\<close> have "path_component S x (pathstart g)"
+    by (simp add: path_connected_component)
+  then obtain f where "path f" and f: "path_image f \<subseteq> S" "pathstart f = x" "pathfinish f = pathstart g"
+    by (auto simp: path_component_def)
+  then have "path (f +++ g)"
+    by (simp add: \<open>arc g\<close> arc_imp_path)
+  then obtain h where "arc h"
+                  and h: "path_image h \<subseteq> path_image (f +++ g)" "pathstart h = x" "pathfinish h = pathfinish g"
+    apply (rule path_contains_arc [of "f +++ g" x "pathfinish g"])
+    using f \<open>x \<notin> V\<close> \<open>pathfinish g \<in> V\<close> by auto
+  have "h ` {0..1} - {h 1} \<subseteq> S"
+    using f g h apply (clarsimp simp: path_image_join)
+    apply (simp add: path_image_def pathfinish_def subset_iff image_def Bex_def)
+    by (metis le_less)
+  then have "h ` {0..<1} \<subseteq> S"
+    using \<open>arc h\<close> by (force simp: arc_def dest: inj_onD)
+  then show thesis
+    apply (rule that [OF \<open>arc h\<close>])
+    using h \<open>pathfinish g \<in> V\<close> by auto
+qed
+
+lemma dense_access_fp_aux:
+  fixes S :: "'a::{complete_space,real_normed_vector} set"
+  assumes S: "open S" "connected S"
+      and opeSU: "openin (subtopology euclidean (frontier S)) U"
+      and opeSV: "openin (subtopology euclidean (frontier S)) V"
+      and "V \<noteq> {}" "\<not> U \<subseteq> V"
+  obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
+proof -
+  have "S \<noteq> {}"
+    using opeSV \<open>V \<noteq> {}\<close> by (metis frontier_empty openin_subtopology_empty)
+  then obtain x where "x \<in> S" by auto
+  obtain g where "arc g" and g: "g ` {0..<1} \<subseteq> S" "pathstart g = x" "pathfinish g \<in> V"
+    using dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close> \<open>V \<noteq> {}\<close> opeSV] by blast
+  obtain h where "arc h" and h: "h ` {0..<1} \<subseteq> S" "pathstart h = x" "pathfinish h \<in> U - {pathfinish g}"
+  proof (rule dense_accessible_frontier_points_connected [OF S \<open>x \<in> S\<close>])
+    show "U - {pathfinish g} \<noteq> {}"
+      using \<open>pathfinish g \<in> V\<close> \<open>\<not> U \<subseteq> V\<close> by blast
+    show "openin (subtopology euclidean (frontier S)) (U - {pathfinish g})"
+      by (simp add: opeSU openin_delete)
+  qed auto
+  obtain \<gamma> where "arc \<gamma>"
+             and \<gamma>: "path_image \<gamma> \<subseteq> path_image (reversepath h +++ g)"
+                    "pathstart \<gamma> = pathfinish h" "pathfinish \<gamma> = pathfinish g"
+  proof (rule path_contains_arc [of "(reversepath h +++ g)" "pathfinish h" "pathfinish g"])
+    show "path (reversepath h +++ g)"
+      by (simp add: \<open>arc g\<close> \<open>arc h\<close> \<open>pathstart g = x\<close> \<open>pathstart h = x\<close> arc_imp_path)
+    show "pathstart (reversepath h +++ g) = pathfinish h"
+         "pathfinish (reversepath h +++ g) = pathfinish g"
+      by auto
+    show "pathfinish h \<noteq> pathfinish g"
+      using \<open>pathfinish h \<in> U - {pathfinish g}\<close> by auto
+  qed auto
+  show ?thesis
+  proof
+    show "arc \<gamma>" "pathstart \<gamma> \<in> U" "pathfinish \<gamma> \<in> V"
+      using \<gamma> \<open>arc \<gamma>\<close> \<open>pathfinish h \<in> U - {pathfinish g}\<close>  \<open>pathfinish g \<in> V\<close> by auto
+    have "\<gamma> ` {0..1} - {\<gamma> 0, \<gamma> 1} \<subseteq> S"
+      using \<gamma> g h
+      apply (simp add: path_image_join)
+      apply (simp add: path_image_def pathstart_def pathfinish_def subset_iff image_def Bex_def)
+      by (metis linorder_neqE_linordered_idom not_less)
+    then show "\<gamma> ` {0<..<1} \<subseteq> S"
+      using \<open>arc h\<close> \<open>arc \<gamma>\<close>
+      by (metis arc_imp_simple_path path_image_def pathfinish_def pathstart_def simple_path_endless)
+  qed
+qed
+
+lemma dense_accessible_frontier_point_pairs:
+  fixes S :: "'a::{complete_space,real_normed_vector} set"
+  assumes S: "open S" "connected S"
+      and opeSU: "openin (subtopology euclidean (frontier S)) U"
+      and opeSV: "openin (subtopology euclidean (frontier S)) V"
+      and "U \<noteq> {}" "V \<noteq> {}" "U \<noteq> V"
+    obtains g where "arc g" "pathstart g \<in> U" "pathfinish g \<in> V" "g ` {0<..<1} \<subseteq> S"
+proof -
+  consider "\<not> U \<subseteq> V" | "\<not> V \<subseteq> U"
+    using \<open>U \<noteq> V\<close> by blast
+  then show ?thesis
+  proof cases
+    case 1 then show ?thesis
+      using assms dense_access_fp_aux [OF S opeSU opeSV] that by blast
+  next
+    case 2
+    obtain g where "arc g" and g: "pathstart g \<in> V" "pathfinish g \<in> U" "g ` {0<..<1} \<subseteq> S"
+      using assms dense_access_fp_aux [OF S opeSV opeSU] "2" by blast
+    show ?thesis
+    proof
+      show "arc (reversepath g)"
+        by (simp add: \<open>arc g\<close> arc_reversepath)
+      show "pathstart (reversepath g) \<in> U" "pathfinish (reversepath g) \<in> V"
+        using g by auto
+      show "reversepath g ` {0<..<1} \<subseteq> S"
+        using g by (auto simp: reversepath_def)
+    qed
+  qed
+qed
+
+end
--- a/src/HOL/Analysis/Brouwer_Fixpoint.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -2021,6 +2021,55 @@
     done
 qed
 
+lemma connected_sphere_eq:
+  fixes a :: "'a :: euclidean_space"
+  shows "connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
+    (is "?lhs = ?rhs")
+proof (cases r "0::real" rule: linorder_cases)
+  case less
+  then show ?thesis by auto
+next
+  case equal
+  then show ?thesis by auto
+next
+  case greater
+  show ?thesis
+  proof
+    assume L: ?lhs
+    have "False" if 1: "DIM('a) = 1"
+    proof -
+      obtain x y where xy: "sphere a r = {x,y}" "x \<noteq> y"
+        using sphere_1D_doubleton [OF 1 greater]
+        by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist)
+      then have "finite (sphere a r)"
+        by auto
+      with L \<open>r > 0\<close> show "False"
+        apply (auto simp: connected_finite_iff_sing)
+        using xy by auto
+    qed
+    with greater show ?rhs
+      by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq)
+  next
+    assume ?rhs
+    then show ?lhs
+      using connected_sphere greater by auto
+  qed
+qed
+
+lemma path_connected_sphere_eq:
+  fixes a :: "'a :: euclidean_space"
+  shows "path_connected(sphere a r) \<longleftrightarrow> 2 \<le> DIM('a) \<or> r \<le> 0"
+         (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    using connected_sphere_eq path_connected_imp_connected by blast
+next
+  assume R: ?rhs
+  then show ?lhs
+    by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere)
+qed
+
 proposition frontier_subset_retraction:
   fixes S :: "'a::euclidean_space set"
   assumes "bounded S" and fros: "frontier S \<subseteq> T"
@@ -2073,7 +2122,7 @@
   ultimately show False by simp
 qed
 
-subsection\<open>Retractions\<close>
+subsection\<open>More Properties of Retractions\<close>
 
 lemma retraction:
    "retraction s t r \<longleftrightarrow>
@@ -2272,6 +2321,60 @@
   using assms
   by (auto simp: retract_of_def retraction intro!: retraction_imp_quotient_map elim!: locally_path_connected_quotient_image)
 
+subsubsection\<open>A few simple lemmas about deformation retracts\<close>
+
+lemma deformation_retract_imp_homotopy_eqv:
+  fixes S :: "'a::euclidean_space set"
+  assumes "homotopic_with (\<lambda>x. True) S S id r" "retraction S T r"
+    shows "S homotopy_eqv T"
+  apply (simp add: homotopy_eqv_def)
+  apply (rule_tac x=r in exI)
+  using assms apply (simp add: retraction_def)
+  apply (rule_tac x=id in exI)
+  apply (auto simp: continuous_on_id)
+   apply (metis homotopic_with_symD)
+  by (metis continuous_on_id' homotopic_with_equal homotopic_with_symD id_apply image_id subset_refl)
+
+lemma deformation_retract:
+  fixes S :: "'a::euclidean_space set"
+    shows "(\<exists>r. homotopic_with (\<lambda>x. True) S S id r \<and> retraction S T r) \<longleftrightarrow>
+           T retract_of S \<and> (\<exists>f. homotopic_with (\<lambda>x. True) S S id f \<and> f ` S \<subseteq> T)"
+    (is "?lhs = ?rhs")
+proof
+  assume ?lhs
+  then show ?rhs
+    by (auto simp: retract_of_def retraction_def)
+next
+  assume ?rhs
+  then show ?lhs
+    apply (clarsimp simp add: retract_of_def retraction_def)
+    apply (rule_tac x=r in exI, simp)
+     apply (rule homotopic_with_trans, assumption)
+     apply (rule_tac f = "r \<circ> f" and g="r \<circ> id" in homotopic_with_eq)
+        apply (rule_tac Y=S in homotopic_compose_continuous_left)
+         apply (auto simp: homotopic_with_sym)
+    done
+qed
+
+lemma deformation_retract_of_contractible_sing:
+  fixes S :: "'a::euclidean_space set"
+  assumes "contractible S" "a \<in> S"
+  obtains r where "homotopic_with (\<lambda>x. True) S S id r" "retraction S {a} r"
+proof -
+  have "{a} retract_of S"
+    by (simp add: \<open>a \<in> S\<close>)
+  moreover have "homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)"
+    using assms
+    apply (clarsimp simp add: contractible_def)
+    apply (rule homotopic_with_trans, assumption)
+    by (metis assms(1) contractible_imp_path_connected homotopic_constant_maps homotopic_with_sym homotopic_with_trans insert_absorb insert_not_empty path_component_mem(1) path_connected_component)
+  moreover have "(\<lambda>x. a) ` S \<subseteq> {a}"
+    by (simp add: image_subsetI)
+  ultimately show ?thesis
+    using that deformation_retract  by metis
+qed
+
+
 subsection\<open>Punctured affine hulls, etc.\<close>
 
 lemma continuous_on_compact_surface_projection_aux:
@@ -2493,6 +2596,23 @@
 by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def
           rel_frontier_retract_of_punctured_affine_hull)
 
+lemma homotopy_eqv_rel_frontier_punctured_convex:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "bounded S" "a \<in> rel_interior S" "convex T" "rel_frontier S \<subseteq> T" "T \<subseteq> affine hull S"
+  shows "(rel_frontier S) homotopy_eqv (T - {a})"
+  apply (rule rel_frontier_deformation_retract_of_punctured_convex [of S T])
+  using assms
+  apply auto
+  apply (subst homotopy_eqv_sym)
+  using deformation_retract_imp_homotopy_eqv by blast
+
+lemma homotopy_eqv_rel_frontier_punctured_affine_hull:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "bounded S" "a \<in> rel_interior S"
+    shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})"
+apply (rule homotopy_eqv_rel_frontier_punctured_convex)
+  using assms rel_frontier_affine_hull  by force+
+
 lemma path_connected_sphere_gen:
   assumes "convex S" "bounded S" "aff_dim S \<noteq> 1"
   shows "path_connected(rel_frontier S)"
@@ -3931,7 +4051,6 @@
   assumes "AR S" "AR T" shows "AR(S \<times> T)"
 using assms by (simp add: AR_ANR ANR_Times contractible_Times)
 
-
 lemma ENR_rel_frontier_convex:
   fixes S :: "'a::euclidean_space set"
   assumes "bounded S" "convex S"
@@ -3972,6 +4091,297 @@
   assumes "bounded S" "convex S"
     shows "ANR(rel_frontier S)"
 by (simp add: ENR_imp_ANR ENR_rel_frontier_convex assms)
+    
+lemma ENR_closedin_Un_local:
+  fixes S :: "'a::euclidean_space set"
+  shows "\<lbrakk>ENR S; ENR T; ENR(S \<inter> T);
+          closedin (subtopology euclidean (S \<union> T)) S; closedin (subtopology euclidean (S \<union> T)) T\<rbrakk>
+        \<Longrightarrow> ENR(S \<union> T)"
+by (simp add: ENR_ANR ANR_closed_Un_local locally_compact_closedin_Un)
+
+lemma ENR_closed_Un:
+  fixes S :: "'a::euclidean_space set"
+  shows "\<lbrakk>closed S; closed T; ENR S; ENR T; ENR(S \<inter> T)\<rbrakk> \<Longrightarrow> ENR(S \<union> T)"
+by (auto simp: closed_subset ENR_closedin_Un_local)
+
+lemma absolute_retract_Un:
+  fixes S :: "'a::euclidean_space set"
+  shows "\<lbrakk>S retract_of UNIV; T retract_of UNIV; (S \<inter> T) retract_of UNIV\<rbrakk>
+         \<Longrightarrow> (S \<union> T) retract_of UNIV"
+  by (meson AR_closed_Un_local_aux closed_subset retract_of_UNIV retract_of_imp_subset)
+
+lemma retract_from_Un_Int:
+  fixes S :: "'a::euclidean_space set"
+  assumes clS: "closedin (subtopology euclidean (S \<union> T)) S"
+      and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+      and Un: "(S \<union> T) retract_of U" and Int: "(S \<inter> T) retract_of T"
+    shows "S retract_of U"
+proof -
+  obtain r where r: "continuous_on T r" "r ` T \<subseteq> S \<inter> T" "\<forall>x\<in>S \<inter> T. r x = x"
+    using Int by (auto simp: retraction_def retract_of_def)
+  have "S retract_of S \<union> T"
+    unfolding retraction_def retract_of_def
+  proof (intro exI conjI)
+    show "continuous_on (S \<union> T) (\<lambda>x. if x \<in> S then x else r x)"
+      apply (rule continuous_on_cases_local [OF clS clT])
+      using r by (auto simp: continuous_on_id)
+  qed (use r in auto)
+  also have "... retract_of U"
+    by (rule Un)
+  finally show ?thesis .
+qed
+
+lemma AR_from_Un_Int_local:
+  fixes S :: "'a::euclidean_space set"
+  assumes clS: "closedin (subtopology euclidean (S \<union> T)) S"
+      and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+      and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)"
+    shows "AR S"
+  apply (rule AR_retract_of_AR [OF Un])
+  by (meson AR_imp_retract clS clT closedin_closed_subset local.Int retract_from_Un_Int retract_of_refl sup_ge2)
+
+lemma AR_from_Un_Int_local':
+  fixes S :: "'a::euclidean_space set"
+  assumes "closedin (subtopology euclidean (S \<union> T)) S"
+      and "closedin (subtopology euclidean (S \<union> T)) T"
+      and "AR(S \<union> T)" "AR(S \<inter> T)"
+    shows "AR T"
+  using AR_from_Un_Int_local [of T S] assms by (simp add: Un_commute Int_commute)
+
+lemma AR_from_Un_Int:
+  fixes S :: "'a::euclidean_space set"
+  assumes clo: "closed S" "closed T" and Un: "AR(S \<union> T)" and Int: "AR(S \<inter> T)"
+  shows "AR S"
+  by (metis AR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest)
+
+lemma ANR_from_Un_Int_local:
+  fixes S :: "'a::euclidean_space set"
+  assumes clS: "closedin (subtopology euclidean (S \<union> T)) S"
+      and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+      and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)"
+    shows "ANR S"
+proof -
+  obtain V where clo: "closedin (subtopology euclidean (S \<union> T)) (S \<inter> T)"
+             and ope: "openin (subtopology euclidean (S \<union> T)) V"
+             and ret: "S \<inter> T retract_of V"
+    using ANR_imp_neighbourhood_retract [OF Int] by (metis clS clT closedin_Int)
+  then obtain r where r: "continuous_on V r" and rim: "r ` V \<subseteq> S \<inter> T" and req: "\<forall>x\<in>S \<inter> T. r x = x"
+    by (auto simp: retraction_def retract_of_def)
+  have Vsub: "V \<subseteq> S \<union> T"
+    by (meson ope openin_contains_cball)
+  have Vsup: "S \<inter> T \<subseteq> V"
+    by (simp add: retract_of_imp_subset ret)
+  then have eq: "S \<union> V = ((S \<union> T) - T) \<union> V"
+    by auto
+  have eq': "S \<union> V = S \<union> (V \<inter> T)"
+    using Vsub by blast
+  have "continuous_on (S \<union> V \<inter> T) (\<lambda>x. if x \<in> S then x else r x)"
+  proof (rule continuous_on_cases_local)
+    show "closedin (subtopology euclidean (S \<union> V \<inter> T)) S"
+      using clS closedin_subset_trans inf.boundedE by blast
+    show "closedin (subtopology euclidean (S \<union> V \<inter> T)) (V \<inter> T)"
+      using clT Vsup by (auto simp: closedin_closed)
+    show "continuous_on (V \<inter> T) r"
+      by (meson Int_lower1 continuous_on_subset r)
+  qed (use req continuous_on_id in auto)
+  with rim have "S retract_of S \<union> V"
+    unfolding retraction_def retract_of_def
+    apply (rule_tac x="\<lambda>x. if x \<in> S then x else r x" in exI)
+    apply (auto simp: eq')
+    done
+  then show ?thesis
+    using ANR_neighborhood_retract [OF Un]
+    using \<open>S \<union> V = S \<union> T - T \<union> V\<close> clT ope by fastforce
+qed
+
+lemma ANR_from_Un_Int:
+  fixes S :: "'a::euclidean_space set"
+  assumes clo: "closed S" "closed T" and Un: "ANR(S \<union> T)" and Int: "ANR(S \<inter> T)"
+  shows "ANR S"
+  by (metis ANR_from_Un_Int_local [OF _ _ Un Int] Un_commute clo closed_closedin closedin_closed_subset inf_sup_absorb subtopology_UNIV top_greatest)
+
+proposition ANR_finite_Union_convex_closed:
+  fixes \<T> :: "'a::euclidean_space set set"
+  assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C"
+  shows "ANR(\<Union>\<T>)"
+proof -
+  have "ANR(\<Union>\<T>)" if "card \<T> < n" for n
+  using assms that
+  proof (induction n arbitrary: \<T>)
+    case 0 then show ?case by simp
+  next
+    case (Suc n)
+    have "ANR(\<Union>\<U>)" if "finite \<U>" "\<U> \<subseteq> \<T>" for \<U>
+      using that
+    proof (induction \<U>)
+      case empty
+      then show ?case  by simp
+    next
+      case (insert C \<U>)
+      have "ANR (C \<union> \<Union>\<U>)"
+      proof (rule ANR_closed_Un)
+        show "ANR (C \<inter> \<Union>\<U>)"
+          unfolding Int_Union
+        proof (rule Suc)
+          show "finite (op \<inter> C ` \<U>)"
+            by (simp add: insert.hyps(1))
+          show "\<And>Ca. Ca \<in> op \<inter> C ` \<U> \<Longrightarrow> closed Ca"
+            by (metis (no_types, hide_lams) Suc.prems(2) closed_Int subsetD imageE insert.prems insertI1 insertI2)
+          show "\<And>Ca. Ca \<in> op \<inter> C ` \<U> \<Longrightarrow> convex Ca"
+            by (metis (mono_tags, lifting) Suc.prems(3) convex_Int imageE insert.prems insert_subset subsetCE)
+          show "card (op \<inter> C ` \<U>) < n"
+          proof -
+            have "card \<T> \<le> n"
+              by (meson Suc.prems(4) not_less not_less_eq)
+            then show ?thesis
+              by (metis Suc.prems(1) card_image_le card_seteq insert.hyps insert.prems insert_subset le_trans not_less)
+          qed
+        qed
+        show "closed (\<Union>\<U>)"
+          using Suc.prems(2) insert.hyps(1) insert.prems by blast
+      qed (use Suc.prems convex_imp_ANR insert.prems insert.IH in auto)
+      then show ?case
+        by simp
+    qed
+    then show ?case
+      using Suc.prems(1) by blast
+  qed
+  then show ?thesis
+    by blast
+qed
+
+
+lemma finite_imp_ANR:
+  fixes S :: "'a::euclidean_space set"
+  assumes "finite S"
+  shows "ANR S"
+proof -
+  have "ANR(\<Union>x \<in> S. {x})"
+    by (blast intro: ANR_finite_Union_convex_closed assms)
+  then show ?thesis
+    by simp
+qed
+
+lemma ANR_insert:
+  fixes S :: "'a::euclidean_space set"
+  assumes "ANR S" "closed S"
+  shows "ANR(insert a S)"
+  by (metis ANR_closed_Un ANR_empty ANR_singleton Diff_disjoint Diff_insert_absorb assms closed_singleton insert_absorb insert_is_Un)
+
+lemma ANR_path_component_ANR:
+  fixes S :: "'a::euclidean_space set"
+  shows "ANR S \<Longrightarrow> ANR(path_component_set S x)"
+  using ANR_imp_locally_path_connected ANR_openin openin_path_component_locally_path_connected by blast
+
+lemma ANR_connected_component_ANR:
+  fixes S :: "'a::euclidean_space set"
+  shows "ANR S \<Longrightarrow> ANR(connected_component_set S x)"
+  by (metis ANR_openin openin_connected_component_locally_connected ANR_imp_locally_connected)
+
+lemma ANR_component_ANR:
+  fixes S :: "'a::euclidean_space set"
+  assumes "ANR S" "c \<in> components S"
+  shows "ANR c"
+  by (metis ANR_connected_component_ANR assms componentsE)
+
+subsection\<open>Original ANR material, now for ENRs.\<close>
+
+lemma ENR_bounded:
+  fixes S :: "'a::euclidean_space set"
+  assumes "bounded S"
+  shows "ENR S \<longleftrightarrow> (\<exists>U. open U \<and> bounded U \<and> S retract_of U)"
+         (is "?lhs = ?rhs")
+proof
+  obtain r where "0 < r" and r: "S \<subseteq> ball 0 r"
+    using bounded_subset_ballD assms by blast
+  assume ?lhs
+  then show ?rhs
+    apply (clarsimp simp: ENR_def)
+    apply (rule_tac x="ball 0 r \<inter> U" in exI, auto)
+    using r retract_of_imp_subset retract_of_subset by fastforce
+next
+  assume ?rhs
+  then show ?lhs
+    using ENR_def by blast
+qed
+
+lemma absolute_retract_imp_AR_gen:
+  fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
+  assumes "S retract_of T" "convex T" "T \<noteq> {}" "S homeomorphic S'" "closedin (subtopology euclidean U) S'"
+  shows "S' retract_of U"
+proof -
+  have "AR T"
+    by (simp add: assms convex_imp_AR)
+  then have "AR S"
+    using AR_retract_of_AR assms by auto
+  then show ?thesis
+    using assms AR_imp_absolute_retract by metis
+qed
+
+lemma absolute_retract_imp_AR:
+  fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
+  assumes "S retract_of UNIV" "S homeomorphic S'" "closed S'"
+  shows "S' retract_of UNIV"
+  using AR_imp_absolute_retract_UNIV assms retract_of_UNIV by blast
+
+lemma homeomorphic_compact_arness:
+  fixes S :: "'a::euclidean_space set" and S' :: "'b::euclidean_space set"
+  assumes "S homeomorphic S'"
+  shows "compact S \<and> S retract_of UNIV \<longleftrightarrow> compact S' \<and> S' retract_of UNIV"
+  using assms homeomorphic_compactness
+  apply auto
+   apply (meson assms compact_AR homeomorphic_AR_iff_AR homeomorphic_compactness)+
+  done
+
+lemma absolute_retract_from_Un_Int:
+  fixes S :: "'a::euclidean_space set"
+  assumes "(S \<union> T) retract_of UNIV" "(S \<inter> T) retract_of UNIV" "closed S" "closed T"
+  shows "S retract_of UNIV"
+  using AR_from_Un_Int assms retract_of_UNIV by auto
+
+lemma ENR_from_Un_Int_gen:
+  fixes S :: "'a::euclidean_space set"
+  assumes "closedin (subtopology euclidean (S \<union> T)) S" "closedin (subtopology euclidean (S \<union> T)) T" "ENR(S \<union> T)" "ENR(S \<inter> T)"
+  shows "ENR S"
+  apply (simp add: ENR_ANR)
+  using ANR_from_Un_Int_local ENR_ANR assms locally_compact_closedin by blast
+
+
+lemma ENR_from_Un_Int:
+  fixes S :: "'a::euclidean_space set"
+  assumes "closed S" "closed T" "ENR(S \<union> T)" "ENR(S \<inter> T)"
+  shows "ENR S"
+  by (meson ENR_from_Un_Int_gen assms closed_subset sup_ge1 sup_ge2)
+
+
+lemma ENR_finite_Union_convex_closed:
+  fixes \<T> :: "'a::euclidean_space set set"
+  assumes \<T>: "finite \<T>" and clo: "\<And>C. C \<in> \<T> \<Longrightarrow> closed C" and con: "\<And>C. C \<in> \<T> \<Longrightarrow> convex C"
+  shows "ENR(\<Union> \<T>)"
+  by (simp add: ENR_ANR ANR_finite_Union_convex_closed \<T> clo closed_Union closed_imp_locally_compact con)
+
+lemma finite_imp_ENR:
+  fixes S :: "'a::euclidean_space set"
+  shows "finite S \<Longrightarrow> ENR S"
+  by (simp add: ENR_ANR finite_imp_ANR finite_imp_closed closed_imp_locally_compact)
+
+lemma ENR_insert:
+  fixes S :: "'a::euclidean_space set"
+  assumes "closed S" "ENR S"
+  shows "ENR(insert a S)"
+proof -
+  have "ENR ({a} \<union> S)"
+    by (metis ANR_insert ENR_ANR Un_commute Un_insert_right assms closed_imp_locally_compact closed_insert sup_bot_right)
+  then show ?thesis
+    by auto
+qed
+
+lemma ENR_path_component_ENR:
+  fixes S :: "'a::euclidean_space set"
+  assumes "ENR S"
+  shows "ENR(path_component_set S x)"
+  by (metis ANR_imp_locally_path_connected ENR_empty ENR_imp_ANR ENR_openin assms
+            locally_path_connected_2 openin_subtopology_self path_component_eq_empty)
 
 (*UNUSED
 lemma ENR_Times:
@@ -3983,6 +4393,61 @@
   SIMP_TAC[ENR_ANR; ANR_PCROSS; LOCALLY_COMPACT_PCROSS]);;
 *)
 
+subsection\<open>Finally, spheres are ANRs and ENRs\<close>
+
+lemma absolute_retract_homeomorphic_convex_compact:
+  fixes S :: "'a::euclidean_space set" and U :: "'b::euclidean_space set"
+  assumes "S homeomorphic U" "S \<noteq> {}" "S \<subseteq> T" "convex U" "compact U"
+  shows "S retract_of T"
+  by (metis UNIV_I assms compact_AR convex_imp_AR homeomorphic_AR_iff_AR homeomorphic_compactness homeomorphic_empty(1) retract_of_subset subsetI)
+
+lemma frontier_retract_of_punctured_universe:
+  fixes S :: "'a::euclidean_space set"
+  assumes "convex S" "bounded S" "a \<in> interior S"
+  shows "(frontier S) retract_of (- {a})"
+  using rel_frontier_retract_of_punctured_affine_hull
+  by (metis Compl_eq_Diff_UNIV affine_hull_nonempty_interior assms empty_iff rel_frontier_frontier rel_interior_nonempty_interior)
+
+lemma sphere_retract_of_punctured_universe_gen:
+  fixes a :: "'a::euclidean_space"
+  assumes "b \<in> ball a r"
+  shows  "sphere a r retract_of (- {b})"
+proof -
+  have "frontier (cball a r) retract_of (- {b})"
+    apply (rule frontier_retract_of_punctured_universe)
+    using assms by auto
+  then show ?thesis
+    by simp
+qed
+
+lemma sphere_retract_of_punctured_universe:
+  fixes a :: "'a::euclidean_space"
+  assumes "0 < r"
+  shows "sphere a r retract_of (- {a})"
+  by (simp add: assms sphere_retract_of_punctured_universe_gen)
+
+proposition ENR_sphere:
+  fixes a :: "'a::euclidean_space"
+  shows "ENR(sphere a r)"
+proof (cases "0 < r")
+  case True
+  then have "sphere a r retract_of -{a}"
+    by (simp add: sphere_retract_of_punctured_universe)
+  with open_delete show ?thesis
+    by (auto simp: ENR_def)
+next
+  case False
+  then show ?thesis
+    using finite_imp_ENR
+    by (metis finite_insert infinite_imp_nonempty less_linear sphere_eq_empty sphere_trivial)
+qed
+
+corollary ANR_sphere:
+  fixes a :: "'a::euclidean_space"
+  shows "ANR(sphere a r)"
+  by (simp add: ENR_imp_ANR ENR_sphere)
+
+
 subsection\<open>Borsuk homotopy extension theorem\<close>
 
 text\<open>It's only this late so we can use the concept of retraction,
@@ -4259,6 +4724,78 @@
   then show ?thesis by blast
 qed
 
+lemma homotopic_Borsuk_maps_in_bounded_component:
+  fixes a :: "'a :: euclidean_space"
+  assumes "compact S" and "a \<notin> S"and "b \<notin> S"
+      and boc: "bounded (connected_component_set (- S) a)"
+      and hom: "homotopic_with (\<lambda>x. True) S (sphere 0 1)
+                               (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
+                               (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
+   shows "connected_component (- S) a b"
+proof (rule ccontr)
+  assume notcc: "\<not> connected_component (- S) a b"
+  let ?T = "S \<union> connected_component_set (- S) a"
+  have "\<nexists>g. continuous_on (S \<union> connected_component_set (- S) a) g \<and>
+            g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1 \<and>
+            (\<forall>x\<in>S. g x = (x - a) /\<^sub>R norm (x - a))"
+    by (simp add: \<open>a \<notin> S\<close> componentsI non_extensible_Borsuk_map [OF \<open>compact S\<close> _ boc])
+  moreover obtain g where "continuous_on (S \<union> connected_component_set (- S) a) g"
+                          "g ` (S \<union> connected_component_set (- S) a) \<subseteq> sphere 0 1"
+                          "\<And>x. x \<in> S \<Longrightarrow> g x = (x - a) /\<^sub>R norm (x - a)"
+  proof (rule Borsuk_homotopy_extension_homotopic)
+    show "closedin (subtopology euclidean ?T) S"
+      by (simp add: \<open>compact S\<close> closed_subset compact_imp_closed)
+    show "continuous_on ?T (\<lambda>x. (x - b) /\<^sub>R norm (x - b))"
+      by (simp add: \<open>b \<notin> S\<close> notcc continuous_on_Borsuk_map)
+    show "(\<lambda>x. (x - b) /\<^sub>R norm (x - b)) ` ?T \<subseteq> sphere 0 1"
+      by (simp add: \<open>b \<notin> S\<close> notcc Borsuk_map_into_sphere)
+    show "homotopic_with (\<lambda>x. True) S (sphere 0 1)
+             (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) (\<lambda>x. (x - a) /\<^sub>R norm (x - a))"
+      by (simp add: hom homotopic_with_symD)
+    qed (auto simp: ANR_sphere intro: that)
+  ultimately show False by blast
+qed
+
+
+lemma Borsuk_maps_homotopic_in_connected_component_eq:
+  fixes a :: "'a :: euclidean_space"
+  assumes S: "compact S" "a \<notin> S" "b \<notin> S" and 2: "2 \<le> DIM('a)"
+    shows "(homotopic_with (\<lambda>x. True) S (sphere 0 1)
+                   (\<lambda>x. (x - a) /\<^sub>R norm (x - a))
+                   (\<lambda>x. (x - b) /\<^sub>R norm (x - b)) \<longleftrightarrow>
+           connected_component (- S) a b)"
+         (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+  proof (cases "bounded(connected_component_set (- S) a)")
+    case True
+    show ?thesis
+      by (rule homotopic_Borsuk_maps_in_bounded_component [OF S True L])
+  next
+    case not_bo_a: False
+    show ?thesis
+    proof (cases "bounded(connected_component_set (- S) b)")
+      case True
+      show ?thesis
+        using homotopic_Borsuk_maps_in_bounded_component [OF S]
+        by (simp add: L True assms connected_component_sym homotopic_Borsuk_maps_in_bounded_component homotopic_with_sym)
+    next
+      case False
+      then show ?thesis
+        using cobounded_unique_unbounded_component [of "-S" a b] \<open>compact S\<close> not_bo_a
+        by (auto simp: compact_eq_bounded_closed assms connected_component_eq_eq)
+    qed
+  qed
+next
+  assume R: ?rhs
+  then have "path_component (- S) a b"
+    using assms(1) compact_eq_bounded_closed open_Compl open_path_connected_component_set by fastforce
+  then show ?lhs
+    by (simp add: Borsuk_maps_homotopic_in_path_component)
+qed
+
+
 subsection\<open>The complement of a set and path-connectedness\<close>
 
 text\<open>Complement in dimension N > 1 of set homeomorphic to any interval in
--- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -577,9 +577,6 @@
 
 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}"
 
--- a/src/HOL/Analysis/Complex_Transcendental.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Complex_Transcendental.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -3559,4 +3559,70 @@
   shows "S homotopy_eqv T \<Longrightarrow> simply_connected S \<longleftrightarrow> simply_connected T"
   by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality)
 
+
+subsection\<open>Homeomorphism of simple closed curves to circles\<close>
+
+proposition homeomorphic_simple_path_image_circle:
+  fixes a :: complex and \<gamma> :: "real \<Rightarrow> 'a::t2_space"
+  assumes "simple_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and "0 < r"
+  shows "(path_image \<gamma>) homeomorphic sphere a r"
+proof -
+  have "homotopic_loops (path_image \<gamma>) \<gamma> \<gamma>"
+    by (simp add: assms homotopic_loops_refl simple_path_imp_path)
+  then have hom: "homotopic_with (\<lambda>h. True) (sphere 0 1) (path_image \<gamma>)
+               (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi)))"
+    by (rule homotopic_loops_imp_homotopic_circlemaps)
+  have "\<exists>g. homeomorphism (sphere 0 1) (path_image \<gamma>) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) g"
+  proof (rule homeomorphism_compact)
+    show "continuous_on (sphere 0 1) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi)))"
+      using hom homotopic_with_imp_continuous by blast
+    show "inj_on (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) (sphere 0 1)"
+    proof
+      fix x y
+      assume xy: "x \<in> sphere 0 1" "y \<in> sphere 0 1"
+         and eq: "(\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) x = (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) y"
+      then have "(Arg x / (2*pi)) = (Arg y / (2*pi))"
+      proof -
+        have "(Arg x / (2*pi)) \<in> {0..1}" "(Arg y / (2*pi)) \<in> {0..1}"
+          using Arg_ge_0 Arg_lt_2pi dual_order.strict_iff_order by fastforce+
+        with eq show ?thesis
+          using \<open>simple_path \<gamma>\<close> Arg_lt_2pi unfolding simple_path_def o_def
+          by (metis eq_divide_eq_1 not_less_iff_gr_or_eq)
+      qed
+      with xy show "x = y"
+        by (metis Arg Arg_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere)
+    qed
+    have "\<And>z. cmod z = 1 \<Longrightarrow> \<exists>x\<in>{0..1}. \<gamma> (Arg z / (2*pi)) = \<gamma> x"
+       by (metis Arg_ge_0 Arg_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral)
+     moreover have "\<exists>z\<in>sphere 0 1. \<gamma> x = \<gamma> (Arg z / (2*pi))" if "0 \<le> x" "x \<le> 1" for x
+     proof (cases "x=1")
+       case True
+       then show ?thesis
+         apply (rule_tac x=1 in bexI)
+         apply (metis loop Arg_of_real divide_eq_0_iff of_real_1 pathfinish_def pathstart_def \<open>0 \<le> x\<close>, auto)
+         done
+     next
+       case False
+       then have *: "(Arg (exp (\<i>*(2* of_real pi* of_real x))) / (2*pi)) = x"
+         using that by (auto simp: Arg_exp divide_simps)
+       show ?thesis
+         by (rule_tac x="exp(ii* of_real(2*pi*x))" in bexI) (auto simp: *)
+    qed
+    ultimately show "(\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) ` sphere 0 1 = path_image \<gamma>"
+      by (auto simp: path_image_def image_iff)
+    qed auto
+    then have "path_image \<gamma> homeomorphic sphere (0::complex) 1"
+      using homeomorphic_def homeomorphic_sym by blast
+  also have "... homeomorphic sphere a r"
+    by (simp add: assms homeomorphic_spheres)
+  finally show ?thesis .
+qed
+
+lemma homeomorphic_simple_path_images:
+  fixes \<gamma>1 :: "real \<Rightarrow> 'a::t2_space" and \<gamma>2 :: "real \<Rightarrow> 'b::t2_space"
+  assumes "simple_path \<gamma>1" and loop: "pathfinish \<gamma>1 = pathstart \<gamma>1"
+  assumes "simple_path \<gamma>2" and loop: "pathfinish \<gamma>2 = pathstart \<gamma>2"
+  shows "(path_image \<gamma>1) homeomorphic (path_image \<gamma>2)"
+  by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero)
+
 end
--- a/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Convex_Euclidean_Space.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -1314,21 +1314,6 @@
 lemma fst_snd_linear: "linear (\<lambda>(x,y). x + y)"
   unfolding linear_iff by (simp add: algebra_simps)
 
-lemma scaleR_2:
-  fixes x :: "'a::real_vector"
-  shows "scaleR 2 x = x + x"
-  unfolding one_add_one [symmetric] scaleR_left_distrib by simp
-
-lemma scaleR_half_double [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "(1 / 2) *\<^sub>R (a + a) = a"
-proof -
-  have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a"
-    by (metis scaleR_2 scaleR_scaleR)
-  then show ?thesis
-    by simp
-qed
-
 lemma vector_choose_size:
   assumes "0 \<le> c"
   obtains x :: "'a::{real_normed_vector, perfect_space}" where "norm x = c"
--- a/src/HOL/Analysis/Further_Topology.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Further_Topology.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -1134,11 +1134,6 @@
 
 (*Up to extend_map_affine_to_sphere_cofinite_gen*)
 
-lemma closedin_closed_subset:
- "\<lbrakk>closedin (subtopology euclidean U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk>
-             \<Longrightarrow> closedin (subtopology euclidean T) S"
-  by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE)
-
 lemma extend_map_affine_to_sphere1:
   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::topological_space"
   assumes "finite K" "affine U" and contf: "continuous_on (U - K) f"
@@ -2976,6 +2971,133 @@
     by (simp add: aff_dim_cball)
 qed
 
+lemma simply_connected_sphere_eq:
+  fixes a :: "'a::euclidean_space"
+  shows "simply_connected(sphere a r) \<longleftrightarrow> 3 \<le> DIM('a) \<or> r \<le> 0"  (is "?lhs = ?rhs")
+proof (cases "r \<le> 0")
+  case True
+  have "simply_connected (sphere a r)"
+    apply (rule convex_imp_simply_connected)
+    using True less_eq_real_def by auto
+  with True show ?thesis by auto
+next
+  case False
+  show ?thesis
+  proof
+    assume L: ?lhs
+    have "False" if "DIM('a) = 1 \<or> DIM('a) = 2"
+      using that
+    proof
+      assume "DIM('a) = 1"
+      with L show False
+        using connected_sphere_eq simply_connected_imp_connected
+        by (metis False Suc_1 not_less_eq_eq order_refl)
+    next
+      assume "DIM('a) = 2"
+      then have "sphere a r homeomorphic sphere (0::complex) 1"
+        by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one)
+      then have "simply_connected(sphere (0::complex) 1)"
+        using L homeomorphic_simply_connected_eq by blast
+      then obtain a::complex where "homotopic_with (\<lambda>h. True) (sphere 0 1) (sphere 0 1) id (\<lambda>x. a)"
+        apply (simp add: simply_connected_eq_contractible_circlemap)
+        by (metis continuous_on_id' id_apply image_id subset_refl)
+      then show False
+        using contractible_sphere contractible_def not_one_le_zero by blast
+    qed
+    with False show ?rhs
+      apply simp
+      by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3)
+  next
+    assume ?rhs
+    with False show ?lhs by (simp add: simply_connected_sphere)
+  qed
+qed
+
+
+lemma simply_connected_punctured_universe_eq:
+  fixes a :: "'a::euclidean_space"
+  shows "simply_connected(- {a}) \<longleftrightarrow> 3 \<le> DIM('a)"
+proof -
+  have [simp]: "a \<in> rel_interior (cball a 1)"
+    by (simp add: rel_interior_nonempty_interior)
+  have [simp]: "affine hull cball a 1 - {a} = -{a}"
+    by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one)
+  have "simply_connected(- {a}) \<longleftrightarrow> simply_connected(sphere a 1)"
+    apply (rule sym)
+    apply (rule homotopy_eqv_simple_connectedness)
+    using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto
+    done
+  also have "...  \<longleftrightarrow> 3 \<le> DIM('a)"
+    by (simp add: simply_connected_sphere_eq)
+  finally show ?thesis .
+qed
+
+lemma not_simply_connected_circle:
+  fixes a :: complex
+  shows "0 < r \<Longrightarrow> \<not> simply_connected(sphere a r)"
+by (simp add: simply_connected_sphere_eq)
+
+  
+proposition simply_connected_punctured_convex:
+  fixes a :: "'a::euclidean_space"
+  assumes "convex S" and 3: "3 \<le> aff_dim S"
+    shows "simply_connected(S - {a})"
+proof (cases "a \<in> rel_interior S")
+  case True
+  then obtain e where "a \<in> S" "0 < e" and e: "cball a e \<inter> affine hull S \<subseteq> S"
+    by (auto simp: rel_interior_cball)
+  have con: "convex (cball a e \<inter> affine hull S)"
+    by (simp add: convex_Int)
+  have bo: "bounded (cball a e \<inter> affine hull S)"
+    by (simp add: bounded_Int)
+  have "affine hull S \<inter> interior (cball a e) \<noteq> {}"
+    using \<open>0 < e\<close> \<open>a \<in> S\<close> hull_subset by fastforce
+  then have "3 \<le> aff_dim (affine hull S \<inter> cball a e)"
+    by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull])
+  also have "... = aff_dim (cball a e \<inter> affine hull S)"
+    by (simp add: Int_commute)
+  finally have "3 \<le> aff_dim (cball a e \<inter> affine hull S)" .
+  moreover have "rel_frontier (cball a e \<inter> affine hull S) homotopy_eqv S - {a}"
+  proof (rule homotopy_eqv_rel_frontier_punctured_convex)
+    show "a \<in> rel_interior (cball a e \<inter> affine hull S)"
+      by (meson IntI Int_mono \<open>a \<in> S\<close> \<open>0 < e\<close> e \<open>cball a e \<inter> affine hull S \<subseteq> S\<close> ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball)
+    have "closed (cball a e \<inter> affine hull S)"
+      by blast
+    then show "rel_frontier (cball a e \<inter> affine hull S) \<subseteq> S"
+      apply (simp add: rel_frontier_def)
+      using e by blast
+    show "S \<subseteq> affine hull (cball a e \<inter> affine hull S)"
+      by (metis (no_types, lifting) IntI \<open>a \<in> S\<close> \<open>0 < e\<close> affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI)
+    qed (auto simp: assms con bo)
+  ultimately show ?thesis
+    using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo]
+    by blast
+next
+  case False
+  show ?thesis
+    apply (rule contractible_imp_simply_connected)
+    apply (rule contractible_convex_tweak_boundary_points [OF \<open>convex S\<close>])
+     apply (simp add: False rel_interior_subset subset_Diff_insert)
+    by (meson Diff_subset closure_subset subset_trans)
+qed
+
+corollary simply_connected_punctured_universe:
+  fixes a :: "'a::euclidean_space"
+  assumes "3 \<le> DIM('a)"
+  shows "simply_connected(- {a})"
+proof -
+  have [simp]: "affine hull cball a 1 = UNIV"
+    apply auto
+    by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq)
+  have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})"
+    apply (rule homotopy_eqv_simple_connectedness)
+    apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull)
+      apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+
+    done
+  then show ?thesis
+    using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV)
+qed
+
 
 subsection\<open>The power, squaring and exponential functions as covering maps\<close>
 
--- a/src/HOL/Analysis/Homeomorphism.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Homeomorphism.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -8,6 +8,34 @@
 imports Path_Connected
 begin
 
+lemma homeomorphic_spheres':
+  fixes a ::"'a::euclidean_space" and b ::"'b::euclidean_space"
+  assumes "0 < \<delta>" and dimeq: "DIM('a) = DIM('b)"
+  shows "(sphere a \<delta>) homeomorphic (sphere b \<delta>)"
+proof -
+  obtain f :: "'a\<Rightarrow>'b" and g where "linear f" "linear g"
+     and fg: "\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y"
+    by (blast intro: isomorphisms_UNIV_UNIV [OF dimeq])
+  then have "continuous_on UNIV f" "continuous_on UNIV g"
+    using linear_continuous_on linear_linear by blast+
+  then show ?thesis
+    unfolding homeomorphic_minimal
+    apply(rule_tac x="\<lambda>x. b + f(x - a)" in exI)
+    apply(rule_tac x="\<lambda>x. a + g(x - b)" in exI)
+    using assms
+    apply (force intro: continuous_intros
+                  continuous_on_compose2 [of _ f] continuous_on_compose2 [of _ g] simp: dist_commute dist_norm fg)
+    done
+qed
+
+lemma homeomorphic_spheres_gen:
+    fixes a :: "'a::euclidean_space" and b :: "'b::euclidean_space"
+  assumes "0 < r" "0 < s" "DIM('a::euclidean_space) = DIM('b::euclidean_space)"
+  shows "(sphere a r homeomorphic sphere b s)"
+  apply (rule homeomorphic_trans [OF homeomorphic_spheres homeomorphic_spheres'])
+  using assms  apply auto
+  done
+
 subsection \<open>Homeomorphism of all convex compact sets with nonempty interior\<close>
 
 proposition ray_to_rel_frontier:
@@ -1360,4 +1388,1036 @@
    shows "g1 x = g2 x"
 using covering_space_lift_unique_gen [of c p S] in_components_self assms ex_in_conv by blast
 
+lemma covering_space_locally:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes loc: "locally \<phi> C" and cov: "covering_space C p S"
+      and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
+    shows "locally \<psi> S"
+proof -
+  have "locally \<psi> (p ` C)"
+    apply (rule locally_open_map_image [OF loc])
+    using cov covering_space_imp_continuous apply blast
+    using cov covering_space_imp_surjective covering_space_open_map apply blast
+    by (simp add: pim)
+  then show ?thesis
+    using covering_space_imp_surjective [OF cov] by metis
+qed
+
+
+proposition covering_space_locally_eq:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes cov: "covering_space C p S"
+      and pim: "\<And>T. \<lbrakk>T \<subseteq> C; \<phi> T\<rbrakk> \<Longrightarrow> \<psi>(p ` T)"
+      and qim: "\<And>q U. \<lbrakk>U \<subseteq> S; continuous_on U q; \<psi> U\<rbrakk> \<Longrightarrow> \<phi>(q ` U)"
+    shows "locally \<psi> S \<longleftrightarrow> locally \<phi> C"
+         (is "?lhs = ?rhs")
+proof
+  assume L: ?lhs
+  show ?rhs
+  proof (rule locallyI)
+    fix V x
+    assume V: "openin (subtopology euclidean C) V" and "x \<in> V"
+    have "p x \<in> p ` C"
+      by (metis IntE V \<open>x \<in> V\<close> imageI openin_open)
+    then obtain T \<V> where "p x \<in> T"
+                      and opeT: "openin (subtopology euclidean S) T"
+                      and veq: "\<Union>\<V> = {x \<in> C. p x \<in> T}"
+                      and ope: "\<forall>U\<in>\<V>. openin (subtopology euclidean C) U"
+                      and hom: "\<forall>U\<in>\<V>. \<exists>q. homeomorphism U T p q"
+      using cov unfolding covering_space_def by (blast intro: that)
+    have "x \<in> \<Union>\<V>"
+      using V veq \<open>p x \<in> T\<close> \<open>x \<in> V\<close> openin_imp_subset by fastforce
+    then obtain U where "x \<in> U" "U \<in> \<V>"
+      by blast
+    then obtain q where opeU: "openin (subtopology euclidean C) U" and q: "homeomorphism U T p q"
+      using ope hom by blast
+    with V have "openin (subtopology euclidean C) (U \<inter> V)"
+      by blast
+    then have UV: "openin (subtopology euclidean S) (p ` (U \<inter> V))"
+      using cov covering_space_open_map by blast
+    obtain W W' where opeW: "openin (subtopology euclidean S) W" and "\<psi> W'" "p x \<in> W" "W \<subseteq> W'" and W'sub: "W' \<subseteq> p ` (U \<inter> V)"
+      using locallyE [OF L UV] \<open>x \<in> U\<close> \<open>x \<in> V\<close> by blast
+    then have "W \<subseteq> T"
+      by (metis Int_lower1 q homeomorphism_image1 image_Int_subset order_trans)
+    show "\<exists>U Z. openin (subtopology euclidean C) U \<and>
+                 \<phi> Z \<and> x \<in> U \<and> U \<subseteq> Z \<and> Z \<subseteq> V"
+    proof (intro exI conjI)
+      have "openin (subtopology euclidean T) W"
+        by (meson opeW opeT openin_imp_subset openin_subset_trans \<open>W \<subseteq> T\<close>)
+      then have "openin (subtopology euclidean U) (q ` W)"
+        by (meson homeomorphism_imp_open_map homeomorphism_symD q)
+      then show "openin (subtopology euclidean C) (q ` W)"
+        using opeU openin_trans by blast
+      show "\<phi> (q ` W')"
+        by (metis (mono_tags, lifting) Int_subset_iff UV W'sub \<open>\<psi> W'\<close> continuous_on_subset dual_order.trans homeomorphism_def image_Int_subset openin_imp_subset q qim)
+      show "x \<in> q ` W"
+        by (metis \<open>p x \<in> W\<close> \<open>x \<in> U\<close> homeomorphism_def imageI q)
+      show "q ` W \<subseteq> q ` W'"
+        using \<open>W \<subseteq> W'\<close> by blast
+      have "W' \<subseteq> p ` V"
+        using W'sub by blast
+      then show "q ` W' \<subseteq> V"
+        using W'sub homeomorphism_apply1 [OF q] by auto
+      qed
+  qed
+next
+  assume ?rhs
+  then show ?lhs
+    using cov covering_space_locally pim by blast
+qed
+
+lemma covering_space_locally_compact_eq:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "covering_space C p S"
+  shows "locally compact S \<longleftrightarrow> locally compact C"
+  apply (rule covering_space_locally_eq [OF assms])
+   apply (meson assms compact_continuous_image continuous_on_subset covering_space_imp_continuous)
+  using compact_continuous_image by blast
+
+lemma covering_space_locally_connected_eq:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "covering_space C p S"
+    shows "locally connected S \<longleftrightarrow> locally connected C"
+  apply (rule covering_space_locally_eq [OF assms])
+   apply (meson connected_continuous_image assms continuous_on_subset covering_space_imp_continuous)
+  using connected_continuous_image by blast
+
+lemma covering_space_locally_path_connected_eq:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "covering_space C p S"
+    shows "locally path_connected S \<longleftrightarrow> locally path_connected C"
+  apply (rule covering_space_locally_eq [OF assms])
+   apply (meson path_connected_continuous_image assms continuous_on_subset covering_space_imp_continuous)
+  using path_connected_continuous_image by blast
+
+
+lemma covering_space_locally_compact:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "locally compact C" "covering_space C p S"
+  shows "locally compact S"
+  using assms covering_space_locally_compact_eq by blast
+
+
+lemma covering_space_locally_connected:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "locally connected C" "covering_space C p S"
+  shows "locally connected S"
+  using assms covering_space_locally_connected_eq by blast
+
+lemma covering_space_locally_path_connected:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes "locally path_connected C" "covering_space C p S"
+  shows "locally path_connected S"
+  using assms covering_space_locally_path_connected_eq by blast
+
+proposition covering_space_lift_homotopy:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and h :: "real \<times> 'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S"
+      and conth: "continuous_on ({0..1} \<times> U) h"
+      and him: "h ` ({0..1} \<times> U) \<subseteq> S"
+      and heq: "\<And>y. y \<in> U \<Longrightarrow> h (0,y) = p(f y)"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C"
+    obtains k where "continuous_on ({0..1} \<times> U) k"
+                    "k ` ({0..1} \<times> U) \<subseteq> C"
+                    "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
+                    "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
+proof -
+  have "\<exists>V k. openin (subtopology euclidean U) V \<and> y \<in> V \<and>
+              continuous_on ({0..1} \<times> V) k \<and> k ` ({0..1} \<times> V) \<subseteq> C \<and>
+              (\<forall>z \<in> V. k(0, z) = f z) \<and> (\<forall>z \<in> {0..1} \<times> V. h z = p(k z))"
+        if "y \<in> U" for y
+  proof -
+    obtain UU where UU: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (subtopology euclidean S) (UU s) \<and>
+                                        (\<exists>\<V>. \<Union>\<V> = {x. x \<in> C \<and> p x \<in> (UU s)} \<and>
+                                            (\<forall>U \<in> \<V>. openin (subtopology euclidean C) U) \<and>
+                                            pairwise disjnt \<V> \<and>
+                                            (\<forall>U \<in> \<V>. \<exists>q. homeomorphism U (UU s) p q))"
+      using cov unfolding covering_space_def by (metis (mono_tags))
+    then have ope: "\<And>s. s \<in> S \<Longrightarrow> s \<in> (UU s) \<and> openin (subtopology euclidean S) (UU s)"
+      by blast
+    have "\<exists>k n i. open k \<and> open n \<and>
+                  t \<in> k \<and> y \<in> n \<and> i \<in> S \<and> h ` (({0..1} \<inter> k) \<times> (U \<inter> n)) \<subseteq> UU i" if "t \<in> {0..1}" for t
+    proof -
+      have hinS: "h (t, y) \<in> S"
+        using \<open>y \<in> U\<close> him that by blast
+      then have "(t,y) \<in> {z \<in> {0..1} \<times> U. h z \<in> UU (h (t, y))}"
+        using \<open>y \<in> U\<close> \<open>t \<in> {0..1}\<close>  by (auto simp: ope)
+      moreover have ope_01U: "openin (subtopology euclidean ({0..1} \<times> U)) {z. z \<in> ({0..1} \<times> U) \<and> h z \<in> UU(h(t, y))}"
+        using hinS ope continuous_on_open_gen [OF him] conth by blast
+      ultimately obtain V W where opeV: "open V" and "t \<in> {0..1} \<inter> V" "t \<in> {0..1} \<inter> V"
+                              and opeW: "open W" and "y \<in> U" "y \<in> W"
+                              and VW: "({0..1} \<inter> V) \<times> (U \<inter> W)  \<subseteq> {z. z \<in> ({0..1} \<times> U) \<and> h z \<in> UU(h(t, y))}"
+        by (rule Times_in_interior_subtopology) (auto simp: openin_open)
+      then show ?thesis
+        using hinS by blast
+    qed
+    then obtain K NN X where
+              K: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (K t)"
+          and NN: "\<And>t. t \<in> {0..1} \<Longrightarrow> open (NN t)"
+          and inUS: "\<And>t. t \<in> {0..1} \<Longrightarrow> t \<in> K t \<and> y \<in> NN t \<and> X t \<in> S"
+          and him: "\<And>t. t \<in> {0..1} \<Longrightarrow> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t)) \<subseteq> UU (X t)"
+    by (metis (mono_tags))
+    obtain \<T> where "\<T> \<subseteq> ((\<lambda>i. K i \<times> NN i)) ` {0..1}" "finite \<T>" "{0::real..1} \<times> {y} \<subseteq> \<Union>\<T>"
+    proof (rule compactE)
+      show "compact ({0::real..1} \<times> {y})"
+        by (simp add: compact_Times)
+      show "{0..1} \<times> {y} \<subseteq> (\<Union>i\<in>{0..1}. K i \<times> NN i)"
+        using K inUS by auto
+      show "\<And>B. B \<in> (\<lambda>i. K i \<times> NN i) ` {0..1} \<Longrightarrow> open B"
+        using K NN by (auto simp: open_Times)
+    qed blast
+    then obtain tk where "tk \<subseteq> {0..1}" "finite tk"
+                     and tk: "{0::real..1} \<times> {y} \<subseteq> (\<Union>i \<in> tk. K i \<times> NN i)"
+      by (metis (no_types, lifting) finite_subset_image)
+    then have "tk \<noteq> {}"
+      by auto
+    define n where "n = INTER tk NN"
+    have "y \<in> n" "open n"
+      using inUS NN \<open>tk \<subseteq> {0..1}\<close> \<open>finite tk\<close>
+      by (auto simp: n_def open_INT subset_iff)
+    obtain \<delta> where "0 < \<delta>" and \<delta>: "\<And>T. \<lbrakk>T \<subseteq> {0..1}; diameter T < \<delta>\<rbrakk> \<Longrightarrow> \<exists>B\<in>K ` tk. T \<subseteq> B"
+    proof (rule Lebesgue_number_lemma [of "{0..1}" "K ` tk"])
+      show "K ` tk \<noteq> {}"
+        using \<open>tk \<noteq> {}\<close> by auto
+      show "{0..1} \<subseteq> UNION tk K"
+        using tk by auto
+      show "\<And>B. B \<in> K ` tk \<Longrightarrow> open B"
+        using \<open>tk \<subseteq> {0..1}\<close> K by auto
+    qed auto
+    obtain N::nat where N: "N > 1 / \<delta>"
+      using reals_Archimedean2 by blast
+    then have "N > 0"
+      using \<open>0 < \<delta>\<close> order.asym by force
+    have *: "\<exists>V k. openin (subtopology euclidean U) V \<and> y \<in> V \<and>
+                   continuous_on ({0..of_nat n / N} \<times> V) k \<and>
+                   k ` ({0..of_nat n / N} \<times> V) \<subseteq> C \<and>
+                   (\<forall>z\<in>V. k (0, z) = f z) \<and>
+                   (\<forall>z\<in>{0..of_nat n / N} \<times> V. h z = p (k z))" if "n \<le> N" for n
+      using that
+    proof (induction n)
+      case 0
+      show ?case
+        apply (rule_tac x=U in exI)
+        apply (rule_tac x="f \<circ> snd" in exI)
+        apply (intro conjI \<open>y \<in> U\<close> continuous_intros continuous_on_subset [OF contf])
+        using fim  apply (auto simp: heq)
+        done
+    next
+      case (Suc n)
+      then obtain V k where opeUV: "openin (subtopology euclidean U) V"
+                        and "y \<in> V"
+                        and contk: "continuous_on ({0..real n / real N} \<times> V) k"
+                        and kim: "k ` ({0..real n / real N} \<times> V) \<subseteq> C"
+                        and keq: "\<And>z. z \<in> V \<Longrightarrow> k (0, z) = f z"
+                        and heq: "\<And>z. z \<in> {0..real n / real N} \<times> V \<Longrightarrow> h z = p (k z)"
+        using Suc_leD by auto
+      have "n \<le> N"
+        using Suc.prems by auto
+      obtain t where "t \<in> tk" and t: "{real n / real N .. (1 + real n) / real N} \<subseteq> K t"
+      proof (rule bexE [OF \<delta>])
+        show "{real n / real N .. (1 + real n) / real N} \<subseteq> {0..1}"
+          using Suc.prems by (auto simp: divide_simps algebra_simps)
+        show diameter_less: "diameter {real n / real N .. (1 + real n) / real N} < \<delta>"
+          using \<open>0 < \<delta>\<close> N by (auto simp: divide_simps algebra_simps)
+      qed blast
+      have t01: "t \<in> {0..1}"
+        using \<open>t \<in> tk\<close> \<open>tk \<subseteq> {0..1}\<close> by blast
+      obtain \<V> where "\<Union>\<V> = {x. x \<in> C \<and> p x \<in> (UU (X t))}"
+                 and opeC: "\<And>U. U \<in> \<V> \<Longrightarrow> openin (subtopology euclidean C) U"
+                 and "pairwise disjnt \<V>"
+                 and homuu: "\<And>U. U \<in> \<V> \<Longrightarrow> \<exists>q. homeomorphism U (UU (X t)) p q"
+        using inUS [OF t01] UU by meson
+      have n_div_N_in: "real n / real N \<in> {real n / real N .. (1 + real n) / real N}"
+        using N by (auto simp: divide_simps algebra_simps)
+      with t have nN_in_kkt: "real n / real N \<in> K t"
+        by blast
+      have "k (real n / real N, y) \<in> {x. x \<in> C \<and> p x \<in> (UU (X t))}"
+      proof (simp, rule conjI)
+        show "k (real n / real N, y) \<in> C"
+          using \<open>y \<in> V\<close> kim keq by force
+        have "p (k (real n / real N, y)) = h (real n / real N, y)"
+          by (simp add: \<open>y \<in> V\<close> heq)
+        also have "... \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
+          apply (rule imageI)
+           using \<open>y \<in> V\<close> t01 \<open>n \<le> N\<close>
+          apply (simp add: nN_in_kkt \<open>y \<in> U\<close> inUS divide_simps)
+          done
+        also have "... \<subseteq> UU (X t)"
+          using him t01 by blast
+        finally show "p (k (real n / real N, y)) \<in> UU (X t)" .
+      qed
+      then have "k (real n / real N, y) \<in> \<Union>\<V>"
+        using \<open>\<Union>\<V> = {x \<in> C. p x \<in> UU (X t)}\<close> by blast
+      then obtain W where W: "k (real n / real N, y) \<in> W" and "W \<in> \<V>"
+        by blast
+      then obtain p' where opeC': "openin (subtopology euclidean C) W"
+                       and hom': "homeomorphism W (UU (X t)) p p'"
+        using homuu opeC by blast
+      then have "W \<subseteq> C"
+        using openin_imp_subset by blast
+      define W' where "W' = UU(X t)"
+      have opeVW: "openin (subtopology euclidean V) {z \<in> V. (k \<circ> Pair (real n / real N)) z \<in> W}"
+        apply (rule continuous_openin_preimage [OF _ _ opeC'])
+         apply (intro continuous_intros continuous_on_subset [OF contk])
+        using kim apply (auto simp: \<open>y \<in> V\<close> W)
+        done
+      obtain N' where opeUN': "openin (subtopology euclidean U) N'"
+                  and "y \<in> N'" and kimw: "k ` ({(real n / real N)} \<times> N') \<subseteq> W"
+        apply (rule_tac N' = "{z \<in> V. (k \<circ> Pair ((real n / real N))) z \<in> W}" in that)
+        apply (fastforce simp:  \<open>y \<in> V\<close> W intro!: openin_trans [OF opeVW opeUV])+
+        done
+      obtain Q Q' where opeUQ: "openin (subtopology euclidean U) Q"
+                    and cloUQ': "closedin (subtopology euclidean U) Q'"
+                    and "y \<in> Q" "Q \<subseteq> Q'"
+                    and Q': "Q' \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V"
+      proof -
+        obtain VO VX where "open VO" "open VX" and VO: "V = U \<inter> VO" and VX: "N' = U \<inter> VX"
+          using opeUV opeUN' by (auto simp: openin_open)
+        then have "open (NN(t) \<inter> VO \<inter> VX)"
+          using NN t01 by blast
+        then obtain e where "e > 0" and e: "cball y e \<subseteq> NN(t) \<inter> VO \<inter> VX"
+          by (metis Int_iff \<open>N' = U \<inter> VX\<close> \<open>V = U \<inter> VO\<close> \<open>y \<in> N'\<close> \<open>y \<in> V\<close> inUS open_contains_cball t01)
+        show ?thesis
+        proof
+          show "openin (subtopology euclidean U) (U \<inter> ball y e)"
+            by blast
+          show "closedin (subtopology euclidean U) (U \<inter> cball y e)"
+            using e by (auto simp: closedin_closed)
+        qed (use \<open>y \<in> U\<close> \<open>e > 0\<close> VO VX e in auto)
+      qed
+      then have "y \<in> Q'" "Q \<subseteq> (U \<inter> NN(t)) \<inter> N' \<inter> V"
+        by blast+
+      have neq: "{0..real n / real N} \<union> {real n / real N..(1 + real n) / real N} = {0..(1 + real n) / real N}"
+        apply (auto simp: divide_simps)
+        by (metis mult_zero_left of_nat_0_le_iff of_nat_0_less_iff order_trans real_mult_le_cancel_iff1)
+      then have neqQ': "{0..real n / real N} \<times> Q' \<union> {real n / real N..(1 + real n) / real N} \<times> Q' = {0..(1 + real n) / real N} \<times> Q'"
+        by blast
+      have cont: "continuous_on ({0..(1 + real n) / real N} \<times> Q')
+        (\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
+        unfolding neqQ' [symmetric]
+      proof (rule continuous_on_cases_local, simp_all add: neqQ' del: comp_apply)
+        show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \<times> Q'))
+                       ({0..real n / real N} \<times> Q')"
+          apply (simp add: closedin_closed)
+          apply (rule_tac x="{0 .. real n / real N} \<times> UNIV" in exI)
+          using n_div_N_in apply (auto simp: closed_Times)
+          done
+        show "closedin (subtopology euclidean ({0..(1 + real n) / real N} \<times> Q'))
+                       ({real n / real N..(1 + real n) / real N} \<times> Q')"
+          apply (simp add: closedin_closed)
+          apply (rule_tac x="{real n / real N .. (1 + real n) / real N} \<times> UNIV" in exI)
+          apply (auto simp: closed_Times)
+          by (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
+        show "continuous_on ({0..real n / real N} \<times> Q') k"
+          apply (rule continuous_on_subset [OF contk])
+          using Q' by auto
+        have "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') h"
+        proof (rule continuous_on_subset [OF conth])
+          show "{real n / real N..(1 + real n) / real N} \<times> Q' \<subseteq> {0..1} \<times> U"
+            using \<open>N > 0\<close>
+            apply auto
+              apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
+            using Suc.prems order_trans apply fastforce
+            apply (metis IntE  cloUQ' closedin_closed)
+            done
+        qed
+        moreover have "continuous_on (h ` ({real n / real N..(1 + real n) / real N} \<times> Q')) p'"
+        proof (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom']])
+          have "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
+            apply (rule image_mono)
+            using \<open>0 < \<delta>\<close> \<open>N > 0\<close> Suc.prems apply auto
+              apply (meson divide_nonneg_nonneg of_nat_0_le_iff order_trans)
+            using Suc.prems order_trans apply fastforce
+            using t Q' apply auto
+            done
+          with him show "h ` ({real n / real N..(1 + real n) / real N} \<times> Q') \<subseteq> UU (X t)"
+            using t01 by blast
+        qed
+        ultimately show "continuous_on ({real n / real N..(1 + real n) / real N} \<times> Q') (p' \<circ> h)"
+          by (rule continuous_on_compose)
+        have "k (real n / real N, b) = p' (h (real n / real N, b))" if "b \<in> Q'" for b
+        proof -
+          have "k (real n / real N, b) \<in> W"
+            using that Q' kimw  by force
+          then have "k (real n / real N, b) = p' (p (k (real n / real N, b)))"
+            by (simp add:  homeomorphism_apply1 [OF hom'])
+          then show ?thesis
+            using Q' that by (force simp: heq)
+        qed
+        then show "\<And>x. x \<in> {real n / real N..(1 + real n) / real N} \<times> Q' \<and>
+                  x \<in> {0..real n / real N} \<times> Q' \<Longrightarrow> k x = (p' \<circ> h) x"
+          by auto
+      qed
+      have h_in_UU: "h (x, y) \<in> UU (X t)" if "y \<in> Q" "\<not> x \<le> real n / real N" "0 \<le> x" "x \<le> (1 + real n) / real N" for x y
+      proof -
+        have "x \<le> 1"
+          using Suc.prems that order_trans by force
+        moreover have "x \<in> K t"
+          by (meson atLeastAtMost_iff le_less not_le subset_eq t that)
+        moreover have "y \<in> U"
+          using \<open>y \<in> Q\<close> opeUQ openin_imp_subset by blast
+        moreover have "y \<in> NN t"
+          using Q' \<open>Q \<subseteq> Q'\<close> \<open>y \<in> Q\<close> by auto
+        ultimately have "(x, y) \<in> (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
+          using that by auto
+        then have "h (x, y) \<in> h ` (({0..1} \<inter> K t) \<times> (U \<inter> NN t))"
+          by blast
+        also have "... \<subseteq> UU (X t)"
+          by (metis him t01)
+        finally show ?thesis .
+      qed
+      let ?k = "(\<lambda>x. if x \<in> {0..real n / real N} \<times> Q' then k x else (p' \<circ> h) x)"
+      show ?case
+      proof (intro exI conjI)
+        show "continuous_on ({0..real (Suc n) / real N} \<times> Q) ?k"
+          apply (rule continuous_on_subset [OF cont])
+          using \<open>Q \<subseteq> Q'\<close> by auto
+        have "\<And>a b. \<lbrakk>a \<le> real n / real N; b \<in> Q'; 0 \<le> a\<rbrakk> \<Longrightarrow> k (a, b) \<in> C"
+          using kim Q' by force
+        moreover have "\<And>a b. \<lbrakk>b \<in> Q; 0 \<le> a; a \<le> (1 + real n) / real N; \<not> a \<le> real n / real N\<rbrakk> \<Longrightarrow> p' (h (a, b)) \<in> C"
+          apply (rule \<open>W \<subseteq> C\<close> [THEN subsetD])
+          using homeomorphism_image2 [OF hom', symmetric]  h_in_UU  Q' \<open>Q \<subseteq> Q'\<close> \<open>W \<subseteq> C\<close>
+          apply auto
+          done
+        ultimately show "?k ` ({0..real (Suc n) / real N} \<times> Q) \<subseteq> C"
+          using Q' \<open>Q \<subseteq> Q'\<close> by force
+        show "\<forall>z\<in>Q. ?k (0, z) = f z"
+          using Q' keq  \<open>Q \<subseteq> Q'\<close> by auto
+        show "\<forall>z \<in> {0..real (Suc n) / real N} \<times> Q. h z = p(?k z)"
+          using \<open>Q \<subseteq> U \<inter> NN t \<inter> N' \<inter> V\<close> heq apply clarsimp
+          using h_in_UU Q' \<open>Q \<subseteq> Q'\<close> apply (auto simp: homeomorphism_apply2 [OF hom', symmetric])
+          done
+        qed (auto simp: \<open>y \<in> Q\<close> opeUQ)
+    qed
+    show ?thesis
+      using*[OF order_refl] N \<open>0 < \<delta>\<close> by (simp add: split: if_split_asm)
+  qed
+  then obtain V fs where opeV: "\<And>y. y \<in> U \<Longrightarrow> openin (subtopology euclidean U) (V y)"
+          and V: "\<And>y. y \<in> U \<Longrightarrow> y \<in> V y"
+          and contfs: "\<And>y. y \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V y) (fs y)"
+          and *: "\<And>y. y \<in> U \<Longrightarrow> (fs y) ` ({0..1} \<times> V y) \<subseteq> C \<and>
+                            (\<forall>z \<in> V y. fs y (0, z) = f z) \<and>
+                            (\<forall>z \<in> {0..1} \<times> V y. h z = p(fs y z))"
+    by (metis (mono_tags))
+  then have VU: "\<And>y. y \<in> U \<Longrightarrow> V y \<subseteq> U"
+    by (meson openin_imp_subset)
+  obtain k where contk: "continuous_on ({0..1} \<times> U) k"
+             and k: "\<And>x i. \<lbrakk>i \<in> U; x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i\<rbrakk> \<Longrightarrow> k x = fs i x"
+  proof (rule pasting_lemma_exists)
+    show "{0..1} \<times> U \<subseteq> (\<Union>i\<in>U. {0..1} \<times> V i)"
+      apply auto
+      using V by blast
+    show "\<And>i. i \<in> U \<Longrightarrow> openin (subtopology euclidean ({0..1} \<times> U)) ({0..1} \<times> V i)"
+      by (simp add: opeV openin_Times)
+    show "\<And>i. i \<in> U \<Longrightarrow> continuous_on ({0..1} \<times> V i) (fs i)"
+      using contfs by blast
+    show "fs i x = fs j x"  if "i \<in> U" "j \<in> U" and x: "x \<in> {0..1} \<times> U \<inter> {0..1} \<times> V i \<inter> {0..1} \<times> V j"
+         for i j x
+    proof -
+      obtain u y where "x = (u, y)" "y \<in> V i" "y \<in> V j" "0 \<le> u" "u \<le> 1"
+        using x by auto
+      show ?thesis
+      proof (rule covering_space_lift_unique [OF cov, of _ "(0,y)" _ "{0..1} \<times> {y}" h])
+        show "fs i (0, y) = fs j (0, y)"
+          using*V by (simp add: \<open>y \<in> V i\<close> \<open>y \<in> V j\<close> that)
+        show conth_y: "continuous_on ({0..1} \<times> {y}) h"
+          apply (rule continuous_on_subset [OF conth])
+          using VU \<open>y \<in> V j\<close> that by auto
+        show "h ` ({0..1} \<times> {y}) \<subseteq> S"
+          using \<open>y \<in> V i\<close> assms(3) VU that by fastforce
+        show "continuous_on ({0..1} \<times> {y}) (fs i)"
+          using continuous_on_subset [OF contfs] \<open>i \<in> U\<close>
+          by (simp add: \<open>y \<in> V i\<close> subset_iff)
+        show "fs i ` ({0..1} \<times> {y}) \<subseteq> C"
+          using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by fastforce
+        show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs i x)"
+          using "*" \<open>y \<in> V i\<close> \<open>i \<in> U\<close> by blast
+        show "continuous_on ({0..1} \<times> {y}) (fs j)"
+          using continuous_on_subset [OF contfs] \<open>j \<in> U\<close>
+          by (simp add: \<open>y \<in> V j\<close> subset_iff)
+        show "fs j ` ({0..1} \<times> {y}) \<subseteq> C"
+          using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by fastforce
+        show "\<And>x. x \<in> {0..1} \<times> {y} \<Longrightarrow> h x = p (fs j x)"
+          using "*" \<open>y \<in> V j\<close> \<open>j \<in> U\<close> by blast
+        show "connected ({0..1::real} \<times> {y})"
+          using connected_Icc connected_Times connected_sing by blast
+        show "(0, y) \<in> {0..1::real} \<times> {y}"
+          by force
+        show "x \<in> {0..1} \<times> {y}"
+          using \<open>x = (u, y)\<close> x by blast
+      qed
+    qed
+  qed blast
+  show ?thesis
+  proof
+    show "k ` ({0..1} \<times> U) \<subseteq> C"
+      using V*k VU by fastforce
+    show "\<And>y. y \<in> U \<Longrightarrow> k (0, y) = f y"
+      by (simp add: V*k)
+    show "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p (k z)"
+      using V*k by auto
+  qed (auto simp: contk)
+qed
+
+corollary covering_space_lift_homotopy_alt:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and h :: "'c::real_normed_vector \<times> real \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S"
+      and conth: "continuous_on (U \<times> {0..1}) h"
+      and him: "h ` (U \<times> {0..1}) \<subseteq> S"
+      and heq: "\<And>y. y \<in> U \<Longrightarrow> h (y,0) = p(f y)"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> C"
+  obtains k where "continuous_on (U \<times> {0..1}) k"
+                  "k ` (U \<times> {0..1}) \<subseteq> C"
+                  "\<And>y. y \<in> U \<Longrightarrow> k(y, 0) = f y"
+                  "\<And>z. z \<in> U \<times> {0..1} \<Longrightarrow> h z = p(k z)"
+proof -
+  have "continuous_on ({0..1} \<times> U) (h \<circ> (\<lambda>z. (snd z, fst z)))"
+    by (intro continuous_intros continuous_on_subset [OF conth]) auto
+  then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
+                  and kim:  "k ` ({0..1} \<times> U) \<subseteq> C"
+                  and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = f y"
+                  and heqp: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> (h \<circ> (\<lambda>z. Pair (snd z) (fst z))) z = p(k z)"
+    apply (rule covering_space_lift_homotopy [OF cov _ _ _ contf fim])
+    using him  by (auto simp: contf heq)
+  show ?thesis
+    apply (rule_tac k="k \<circ> (\<lambda>z. Pair (snd z) (fst z))" in that)
+       apply (intro continuous_intros continuous_on_subset [OF contk])
+    using kim heqp apply (auto simp: k0)
+    done
+qed
+
+corollary covering_space_lift_homotopic_function:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and g:: "'c::real_normed_vector \<Rightarrow> 'a"
+  assumes cov: "covering_space C p S"
+      and contg: "continuous_on U g"
+      and gim: "g ` U \<subseteq> C"
+      and pgeq: "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+      and hom: "homotopic_with (\<lambda>x. True) U S f f'"
+    obtains g' where "continuous_on U g'" "image g' U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g' y) = f' y"
+proof -
+  obtain h where conth: "continuous_on ({0..1::real} \<times> U) h"
+             and him: "h ` ({0..1} \<times> U) \<subseteq> S"
+             and h0:  "\<And>x. h(0, x) = f x"
+             and h1: "\<And>x. h(1, x) = f' x"
+    using hom by (auto simp: homotopic_with_def)
+  have "\<And>y. y \<in> U \<Longrightarrow> h (0, y) = p (g y)"
+    by (simp add: h0 pgeq)
+  then obtain k where contk: "continuous_on ({0..1} \<times> U) k"
+                  and kim: "k ` ({0..1} \<times> U) \<subseteq> C"
+                  and k0: "\<And>y. y \<in> U \<Longrightarrow> k(0, y) = g y"
+                  and heq: "\<And>z. z \<in> {0..1} \<times> U \<Longrightarrow> h z = p(k z)"
+    using covering_space_lift_homotopy [OF cov conth him _ contg gim] by metis
+  show ?thesis
+  proof
+    show "continuous_on U (k \<circ> Pair 1)"
+      by (meson contk atLeastAtMost_iff continuous_on_o_Pair order_refl zero_le_one)
+    show "(k \<circ> Pair 1) ` U \<subseteq> C"
+      using kim by auto
+    show "\<And>y. y \<in> U \<Longrightarrow> p ((k \<circ> Pair 1) y) = f' y"
+      by (auto simp: h1 heq [symmetric])
+  qed
+qed
+
+corollary covering_space_lift_inessential_function:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" and U :: "'c::real_normed_vector set"
+  assumes cov: "covering_space C p S"
+      and hom: "homotopic_with (\<lambda>x. True) U S f (\<lambda>x. a)"
+  obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+proof (cases "U = {}")
+  case True
+  then show ?thesis
+    using that continuous_on_empty by blast
+next
+  case False
+  then obtain b where b: "b \<in> C" "p b = a"
+    using covering_space_imp_surjective [OF cov] homotopic_with_imp_subset2 [OF hom]
+    by auto
+  then have gim: "(\<lambda>y. b) ` U \<subseteq> C"
+    by blast
+  show ?thesis
+    apply (rule covering_space_lift_homotopic_function
+                  [OF cov continuous_on_const gim _ homotopic_with_symD [OF hom]])
+    using b that apply auto
+    done
+qed
+
+subsection\<open> Lifting of general functions to covering space\<close>
+
+proposition covering_space_lift_path_strong:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and f :: "'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S" and "a \<in> C"
+      and "path g" and pag: "path_image g \<subseteq> S" and pas: "pathstart g = p a"
+    obtains h where "path h" "path_image h \<subseteq> C" "pathstart h = a"
+                and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
+proof -
+  obtain k:: "real \<times> 'c \<Rightarrow> 'a"
+    where contk: "continuous_on ({0..1} \<times> {undefined}) k"
+      and kim: "k ` ({0..1} \<times> {undefined}) \<subseteq> C"
+      and k0:  "k (0, undefined) = a"
+      and pk: "\<And>z. z \<in> {0..1} \<times> {undefined} \<Longrightarrow> p(k z) = (g \<circ> fst) z"
+  proof (rule covering_space_lift_homotopy [OF cov, of "{undefined}" "g \<circ> fst"])
+    show "continuous_on ({0..1::real} \<times> {undefined::'c}) (g \<circ> fst)"
+      apply (intro continuous_intros)
+      using \<open>path g\<close> by (simp add: path_def)
+    show "(g \<circ> fst) ` ({0..1} \<times> {undefined}) \<subseteq> S"
+      using pag by (auto simp: path_image_def)
+    show "(g \<circ> fst) (0, y) = p a" if "y \<in> {undefined}" for y::'c
+      by (metis comp_def fst_conv pas pathstart_def)
+  qed (use assms in auto)
+  show ?thesis
+  proof
+    show "path (k \<circ> (\<lambda>t. Pair t undefined))"
+      unfolding path_def
+      by (intro continuous_on_compose continuous_intros continuous_on_subset [OF contk]) auto
+    show "path_image (k \<circ> (\<lambda>t. (t, undefined))) \<subseteq> C"
+      using kim by (auto simp: path_image_def)
+    show "pathstart (k \<circ> (\<lambda>t. (t, undefined))) = a"
+      by (auto simp: pathstart_def k0)
+    show "\<And>t. t \<in> {0..1} \<Longrightarrow> p ((k \<circ> (\<lambda>t. (t, undefined))) t) = g t"
+      by (auto simp: pk)
+  qed
+qed
+
+corollary covering_space_lift_path:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes cov: "covering_space C p S" and "path g" and pig: "path_image g \<subseteq> S"
+  obtains h where "path h" "path_image h \<subseteq> C" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = g t"
+proof -
+  obtain a where "a \<in> C" "pathstart g = p a"
+    by (metis pig cov covering_space_imp_surjective imageE pathstart_in_path_image subsetCE)
+  show ?thesis
+    using covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close> \<open>path g\<close> pig]
+    by (metis \<open>pathstart g = p a\<close> that)
+qed
+
+  
+proposition covering_space_lift_homotopic_paths:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes cov: "covering_space C p S"
+      and "path g1" and pig1: "path_image g1 \<subseteq> S"
+      and "path g2" and pig2: "path_image g2 \<subseteq> S"
+      and hom: "homotopic_paths S g1 g2"
+      and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
+      and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
+      and h1h2: "pathstart h1 = pathstart h2"
+    shows "homotopic_paths C h1 h2"
+proof -
+  obtain h :: "real \<times> real \<Rightarrow> 'b"
+     where conth: "continuous_on ({0..1} \<times> {0..1}) h"
+       and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
+       and h0: "\<And>x. h (0, x) = g1 x" and h1: "\<And>x. h (1, x) = g2 x"
+       and heq0: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 0) = g1 0"
+       and heq1: "\<And>t. t \<in> {0..1} \<Longrightarrow> h (t, 1) = g1 1"
+    using hom by (auto simp: homotopic_paths_def homotopic_with_def pathstart_def pathfinish_def)
+  obtain k where contk: "continuous_on ({0..1} \<times> {0..1}) k"
+             and kim: "k ` ({0..1} \<times> {0..1}) \<subseteq> C"
+             and kh2: "\<And>y. y \<in> {0..1} \<Longrightarrow> k (y, 0) = h2 0"
+             and hpk: "\<And>z. z \<in> {0..1} \<times> {0..1} \<Longrightarrow> h z = p (k z)"
+    apply (rule covering_space_lift_homotopy_alt [OF cov conth him, of "\<lambda>x. h2 0"])
+    using h1h2 ph1 ph2 apply (force simp: heq0 pathstart_def pathfinish_def)
+    using path_image_def pih2 continuous_on_const by fastforce+
+  have contg1: "continuous_on {0..1} g1" and contg2: "continuous_on {0..1} g2"
+    using \<open>path g1\<close> \<open>path g2\<close> path_def by blast+
+  have g1im: "g1 ` {0..1} \<subseteq> S" and g2im: "g2 ` {0..1} \<subseteq> S"
+    using path_image_def pig1 pig2 by auto
+  have conth1: "continuous_on {0..1} h1" and conth2: "continuous_on {0..1} h2"
+    using \<open>path h1\<close> \<open>path h2\<close> path_def by blast+
+  have h1im: "h1 ` {0..1} \<subseteq> C" and h2im: "h2 ` {0..1} \<subseteq> C"
+    using path_image_def pih1 pih2 by auto
+  show ?thesis
+    unfolding homotopic_paths pathstart_def pathfinish_def
+  proof (intro exI conjI ballI)
+    show keqh1: "k(0, x) = h1 x" if "x \<in> {0..1}" for x
+    proof (rule covering_space_lift_unique [OF cov _ contg1 g1im])
+      show "k (0,0) = h1 0"
+        by (metis atLeastAtMost_iff h1h2 kh2 order_refl pathstart_def zero_le_one)
+      show "continuous_on {0..1} (\<lambda>a. k (0, a))"
+        by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
+      show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 x = p (k (0, x))"
+        by (metis atLeastAtMost_iff h0 hpk zero_le_one mem_Sigma_iff order_refl)
+    qed (use conth1 h1im kim that in \<open>auto simp: ph1\<close>)
+    show "k(1, x) = h2 x" if "x \<in> {0..1}" for x
+    proof (rule covering_space_lift_unique [OF cov _ contg2 g2im])
+      show "k (1,0) = h2 0"
+        by (metis atLeastAtMost_iff kh2 order_refl zero_le_one)
+      show "continuous_on {0..1} (\<lambda>a. k (1, a))"
+        by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
+      show "\<And>x. x \<in> {0..1} \<Longrightarrow> g2 x = p (k (1, x))"
+        by (metis atLeastAtMost_iff h1 hpk mem_Sigma_iff order_refl zero_le_one)
+    qed (use conth2 h2im kim that in \<open>auto simp: ph2\<close>)
+    show "\<And>t. t \<in> {0..1} \<Longrightarrow> (k \<circ> Pair t) 0 = h1 0"
+      by (metis comp_apply h1h2 kh2 pathstart_def)
+    show "(k \<circ> Pair t) 1 = h1 1" if "t \<in> {0..1}" for t
+    proof (rule covering_space_lift_unique
+           [OF cov, of "\<lambda>a. (k \<circ> Pair a) 1" 0 "\<lambda>a. h1 1" "{0..1}"  "\<lambda>x. g1 1"])
+      show "(k \<circ> Pair 0) 1 = h1 1"
+        using keqh1 by auto
+      show "continuous_on {0..1} (\<lambda>a. (k \<circ> Pair a) 1)"
+        apply simp
+        by (intro continuous_intros continuous_on_compose2 [OF contk]) auto
+      show "\<And>x. x \<in> {0..1} \<Longrightarrow> g1 1 = p ((k \<circ> Pair x) 1)"
+        using heq1 hpk by auto
+    qed (use contk kim g1im h1im that in \<open>auto simp: ph1 continuous_on_const\<close>)
+  qed (use contk kim in auto)
+qed
+
+
+corollary covering_space_monodromy:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes cov: "covering_space C p S"
+      and "path g1" and pig1: "path_image g1 \<subseteq> S"
+      and "path g2" and pig2: "path_image g2 \<subseteq> S"
+      and hom: "homotopic_paths S g1 g2"
+      and "path h1" and pih1: "path_image h1 \<subseteq> C" and ph1: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h1 t) = g1 t"
+      and "path h2" and pih2: "path_image h2 \<subseteq> C" and ph2: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h2 t) = g2 t"
+      and h1h2: "pathstart h1 = pathstart h2"
+    shows "pathfinish h1 = pathfinish h2"
+  using covering_space_lift_homotopic_paths [OF assms] homotopic_paths_imp_pathfinish by blast
+
+
+corollary covering_space_lift_homotopic_path:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+  assumes cov: "covering_space C p S"
+      and hom: "homotopic_paths S f f'"
+      and "path g" and pig: "path_image g \<subseteq> C"
+      and a: "pathstart g = a" and b: "pathfinish g = b"
+      and pgeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g t) = f t"
+  obtains g' where "path g'" "path_image g' \<subseteq> C"
+                   "pathstart g' = a" "pathfinish g' = b" "\<And>t. t \<in> {0..1} \<Longrightarrow> p(g' t) = f' t"
+proof (rule covering_space_lift_path_strong [OF cov, of a f'])
+  show "a \<in> C"
+    using a pig by auto
+  show "path f'" "path_image f' \<subseteq> S"
+    using hom homotopic_paths_imp_path homotopic_paths_imp_subset by blast+
+  show "pathstart f' = p a"
+    by (metis a atLeastAtMost_iff hom homotopic_paths_imp_pathstart order_refl pathstart_def pgeq zero_le_one)
+qed (metis (mono_tags, lifting) assms cov covering_space_monodromy hom homotopic_paths_imp_path homotopic_paths_imp_subset pgeq pig)
+
+
+proposition covering_space_lift_general:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and f :: "'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S" and "a \<in> C" "z \<in> U"
+      and U: "path_connected U" "locally path_connected U"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
+      and feq: "f z = p a"
+      and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk>
+                     \<Longrightarrow> \<exists>q. path q \<and> path_image q \<subseteq> C \<and>
+                             pathstart q = a \<and> pathfinish q = a \<and>
+                             homotopic_paths S (f \<circ> r) (p \<circ> q)"
+  obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+proof -
+  have *: "\<exists>g h. path g \<and> path_image g \<subseteq> U \<and>
+                 pathstart g = z \<and> pathfinish g = y \<and>
+                 path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and>
+                 (\<forall>t \<in> {0..1}. p(h t) = f(g t))"
+          if "y \<in> U" for y
+  proof -
+    obtain g where "path g" "path_image g \<subseteq> U" and pastg: "pathstart g = z"
+               and pafig: "pathfinish g = y"
+      using U \<open>z \<in> U\<close> \<open>y \<in> U\<close> by (force simp: path_connected_def)
+    obtain h where "path h" "path_image h \<subseteq> C" "pathstart h = a"
+               and "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = (f \<circ> g) t"
+    proof (rule covering_space_lift_path_strong [OF cov \<open>a \<in> C\<close>])
+      show "path (f \<circ> g)"
+        using \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> contf continuous_on_subset path_continuous_image by blast
+      show "path_image (f \<circ> g) \<subseteq> S"
+        by (metis \<open>path_image g \<subseteq> U\<close> fim image_mono path_image_compose subset_trans)
+      show "pathstart (f \<circ> g) = p a"
+        by (simp add: feq pastg pathstart_compose)
+    qed auto
+    then show ?thesis
+      by (metis \<open>path g\<close> \<open>path_image g \<subseteq> U\<close> comp_apply pafig pastg)
+  qed
+  have "\<exists>l. \<forall>g h. path g \<and> path_image g \<subseteq> U \<and> pathstart g = z \<and> pathfinish g = y \<and>
+                  path h \<and> path_image h \<subseteq> C \<and> pathstart h = a \<and>
+                  (\<forall>t \<in> {0..1}. p(h t) = f(g t))  \<longrightarrow> pathfinish h = l" for y
+  proof -
+    have "pathfinish h = pathfinish h'"
+         if g: "path g" "path_image g \<subseteq> U" "pathstart g = z" "pathfinish g = y"
+            and h: "path h" "path_image h \<subseteq> C" "pathstart h = a"
+            and phg: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)"
+            and g': "path g'" "path_image g' \<subseteq> U" "pathstart g' = z" "pathfinish g' = y"
+            and h': "path h'" "path_image h' \<subseteq> C" "pathstart h' = a"
+            and phg': "\<And>t. t \<in> {0..1} \<Longrightarrow> p(h' t) = f(g' t)"
+         for g h g' h'
+    proof -
+      obtain q where "path q" and piq: "path_image q \<subseteq> C" and pastq: "pathstart q = a" and pafiq: "pathfinish q = a"
+                 and homS: "homotopic_paths S (f \<circ> g +++ reversepath g') (p \<circ> q)"
+        using g g' hom [of "g +++ reversepath g'"] by (auto simp:  subset_path_image_join)
+              have papq: "path (p \<circ> q)"
+                using homS homotopic_paths_imp_path by blast
+              have pipq: "path_image (p \<circ> q) \<subseteq> S"
+                using homS homotopic_paths_imp_subset by blast
+      obtain q' where "path q'" "path_image q' \<subseteq> C"
+                and "pathstart q' = pathstart q" "pathfinish q' = pathfinish q"
+                and pq'_eq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p (q' t) = (f \<circ> g +++ reversepath g') t"
+        using covering_space_lift_homotopic_path [OF cov homotopic_paths_sym [OF homS] \<open>path q\<close> piq refl refl]
+        by auto
+      have "q' t = (h \<circ> op *\<^sub>R 2) t" if "0 \<le> t" "t \<le> 1/2" for t
+      proof (rule covering_space_lift_unique [OF cov, of q' 0 "h \<circ> op *\<^sub>R 2" "{0..1/2}" "f \<circ> g \<circ> op *\<^sub>R 2" t])
+        show "q' 0 = (h \<circ> op *\<^sub>R 2) 0"
+          by (metis \<open>pathstart q' = pathstart q\<close> comp_def g h pastq pathstart_def pth_4(2))
+        show "continuous_on {0..1/2} (f \<circ> g \<circ> op *\<^sub>R 2)"
+          apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path g\<close>] continuous_on_subset [OF contf])
+          using g(2) path_image_def by fastforce+
+        show "(f \<circ> g \<circ> op *\<^sub>R 2) ` {0..1/2} \<subseteq> S"
+          using g(2) path_image_def fim by fastforce
+        show "(h \<circ> op *\<^sub>R 2) ` {0..1/2} \<subseteq> C"
+          using h path_image_def by fastforce
+        show "q' ` {0..1/2} \<subseteq> C"
+          using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce
+        show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> op *\<^sub>R 2) x = p (q' x)"
+          by (auto simp: joinpaths_def pq'_eq)
+        show "\<And>x. x \<in> {0..1/2} \<Longrightarrow> (f \<circ> g \<circ> op *\<^sub>R 2) x = p ((h \<circ> op *\<^sub>R 2) x)"
+          by (simp add: phg)
+        show "continuous_on {0..1/2} q'"
+          by (simp add: continuous_on_path \<open>path q'\<close>)
+        show "continuous_on {0..1/2} (h \<circ> op *\<^sub>R 2)"
+          apply (intro continuous_intros continuous_on_compose continuous_on_path [OF \<open>path h\<close>], force)
+          done
+      qed (use that in auto)
+      moreover have "q' t = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) t" if "1/2 < t" "t \<le> 1" for t
+      proof (rule covering_space_lift_unique [OF cov, of q' 1 "reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" "{1/2<..1}" "f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)" t])
+        show "q' 1 = (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) 1"
+          using h' \<open>pathfinish q' = pathfinish q\<close> pafiq
+          by (simp add: pathstart_def pathfinish_def reversepath_def)
+        show "continuous_on {1/2<..1} (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
+          apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path g'\<close> continuous_on_subset [OF contf])
+          using g' apply simp_all
+          by (auto simp: path_image_def reversepath_def)
+        show "(f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> S"
+          using g'(2) path_image_def fim by (auto simp: image_subset_iff path_image_def reversepath_def)
+        show "q' ` {1/2<..1} \<subseteq> C"
+          using \<open>path_image q' \<subseteq> C\<close> path_image_def by fastforce
+        show "(reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) ` {1/2<..1} \<subseteq> C"
+          using h' by (simp add: path_image_def reversepath_def subset_eq)
+        show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow> (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p (q' x)"
+          by (auto simp: joinpaths_def pq'_eq)
+        show "\<And>x. x \<in> {1/2<..1} \<Longrightarrow>
+                  (f \<circ> reversepath g' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x = p ((reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1)) x)"
+          by (simp add: phg' reversepath_def)
+        show "continuous_on {1/2<..1} q'"
+          by (auto intro: continuous_on_path [OF \<open>path q'\<close>])
+        show "continuous_on {1/2<..1} (reversepath h' \<circ> (\<lambda>t. 2 *\<^sub>R t - 1))"
+          apply (intro continuous_intros continuous_on_compose continuous_on_path \<open>path h'\<close>)
+          using h' apply auto
+          done
+      qed (use that in auto)
+      ultimately have "q' t = (h +++ reversepath h') t" if "0 \<le> t" "t \<le> 1" for t
+        using that by (simp add: joinpaths_def)
+      then have "path(h +++ reversepath h')"
+        by (auto intro: path_eq [OF \<open>path q'\<close>])
+      then show ?thesis
+        by (auto simp: \<open>path h\<close> \<open>path h'\<close>)
+    qed
+    then show ?thesis by metis
+  qed
+  then obtain l :: "'c \<Rightarrow> 'a"
+          where l: "\<And>y g h. \<lbrakk>path g; path_image g \<subseteq> U; pathstart g = z; pathfinish g = y;
+                             path h; path_image h \<subseteq> C; pathstart h = a;
+                             \<And>t. t \<in> {0..1} \<Longrightarrow> p(h t) = f(g t)\<rbrakk> \<Longrightarrow> pathfinish h = l y"
+    by metis
+  show ?thesis
+  proof
+    show pleq: "p (l y) = f y" if "y \<in> U" for y
+      using*[OF \<open>y \<in> U\<close>]  by (metis l atLeastAtMost_iff order_refl pathfinish_def zero_le_one)
+    show "l z = a"
+      using l [of "linepath z z" z "linepath a a"] by (auto simp: assms)
+    show LC: "l ` U \<subseteq> C"
+      by (clarify dest!: *) (metis (full_types) l pathfinish_in_path_image subsetCE)
+    have "\<exists>T. openin (subtopology euclidean U) T \<and> y \<in> T \<and> T \<subseteq> {x \<in> U. l x \<in> X}"
+         if X: "openin (subtopology euclidean C) X" and "y \<in> U" "l y \<in> X" for X y
+    proof -
+      have "X \<subseteq> C"
+        using X openin_euclidean_subtopology_iff by blast
+      have "f y \<in> S"
+        using fim \<open>y \<in> U\<close> by blast
+      then obtain W \<V>
+              where WV: "f y \<in> W \<and> openin (subtopology euclidean S) W \<and>
+                         (\<Union>\<V> = {x. x \<in> C \<and> p x \<in> W} \<and>
+                          (\<forall>U \<in> \<V>. openin (subtopology euclidean C) U) \<and>
+                          pairwise disjnt \<V> \<and>
+                          (\<forall>U \<in> \<V>. \<exists>q. homeomorphism U W p q))"
+        using cov by (force simp: covering_space_def)
+      then have "l y \<in> \<Union>\<V>"
+        using \<open>X \<subseteq> C\<close> pleq that by auto
+      then obtain W' where "l y \<in> W'" and "W' \<in> \<V>"
+        by blast
+      with WV obtain p' where opeCW': "openin (subtopology euclidean C) W'"
+                          and homUW': "homeomorphism W' W p p'"
+        by blast
+      then have contp': "continuous_on W p'" and p'im: "p' ` W \<subseteq> W'"
+        using homUW' homeomorphism_image2 homeomorphism_cont2 by fastforce+
+      obtain V where "y \<in> V" "y \<in> U" and fimW: "f ` V \<subseteq> W" "V \<subseteq> U"
+                 and "path_connected V" and opeUV: "openin (subtopology euclidean U) V"
+      proof -
+        have "openin (subtopology euclidean U) {c \<in> U. f c \<in> W}"
+          using WV contf continuous_on_open_gen fim by auto
+        then show ?thesis
+          using U WV
+          apply (auto simp: locally_path_connected)
+          apply (drule_tac x="{x. x \<in> U \<and> f x \<in> W}" in spec)
+          apply (drule_tac x=y in spec)
+          apply (auto simp: \<open>y \<in> U\<close> intro: that)
+          done
+      qed
+      have "W' \<subseteq> C" "W \<subseteq> S"
+        using opeCW' WV openin_imp_subset by auto
+      have p'im: "p' ` W \<subseteq> W'"
+        using homUW' homeomorphism_image2 by fastforce
+      show ?thesis
+      proof (intro exI conjI)
+        have "openin (subtopology euclidean S) {x \<in> W. p' x \<in> W' \<inter> X}"
+        proof (rule openin_trans)
+          show "openin (subtopology euclidean W) {x \<in> W. p' x \<in> W' \<inter> X}"
+            apply (rule continuous_openin_preimage [OF contp' p'im])
+            using X \<open>W' \<subseteq> C\<close> apply (auto simp: openin_open)
+            done
+          show "openin (subtopology euclidean S) W"
+            using WV by blast
+        qed
+        then show "openin (subtopology euclidean U) (V \<inter> {x. x \<in> U \<and> f x \<in> {x. x \<in> W \<and> p' x \<in> W' \<inter> X}})"
+          by (intro openin_Int opeUV continuous_openin_preimage [OF contf fim])
+        have "p' (f y) \<in> X"
+          using \<open>l y \<in> W'\<close> homeomorphism_apply1 [OF homUW'] pleq \<open>y \<in> U\<close> \<open>l y \<in> X\<close> by fastforce
+        then show "y \<in> V \<inter> {x \<in> U. f x \<in> {x \<in> W. p' x \<in> W' \<inter> X}}"
+          using \<open>y \<in> U\<close> \<open>y \<in> V\<close> WV p'im by auto
+        show "V \<inter> {x \<in> U. f x \<in> {x \<in> W. p' x \<in> W' \<inter> X}} \<subseteq> {x \<in> U. l x \<in> X}"
+        proof clarsimp
+          fix y'
+          assume y': "y' \<in> V" "y' \<in> U" "f y' \<in> W" "p' (f y') \<in> W'" "p' (f y') \<in> X"
+          then obtain \<gamma> where "path \<gamma>" "path_image \<gamma> \<subseteq> V" "pathstart \<gamma> = y" "pathfinish \<gamma> = y'"
+            by (meson \<open>path_connected V\<close> \<open>y \<in> V\<close> path_connected_def)
+          obtain pp qq where "path pp" "path_image pp \<subseteq> U"
+                             "pathstart pp = z" "pathfinish pp = y"
+                             "path qq" "path_image qq \<subseteq> C" "pathstart qq = a"
+                         and pqqeq: "\<And>t. t \<in> {0..1} \<Longrightarrow> p(qq t) = f(pp t)"
+            using*[OF \<open>y \<in> U\<close>] by blast
+          have finW: "\<And>x. \<lbrakk>0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> f (\<gamma> x) \<in> W"
+            using \<open>path_image \<gamma> \<subseteq> V\<close> by (auto simp: image_subset_iff path_image_def fimW [THEN subsetD])
+          have "pathfinish (qq +++ (p' \<circ> f \<circ> \<gamma>)) = l y'"
+          proof (rule l [of "pp +++ \<gamma>" y' "qq +++ (p' \<circ> f \<circ> \<gamma>)" ])
+            show "path (pp +++ \<gamma>)"
+              by (simp add: \<open>path \<gamma>\<close> \<open>path pp\<close> \<open>pathfinish pp = y\<close> \<open>pathstart \<gamma> = y\<close>)
+            show "path_image (pp +++ \<gamma>) \<subseteq> U"
+              using \<open>V \<subseteq> U\<close> \<open>path_image \<gamma> \<subseteq> V\<close> \<open>path_image pp \<subseteq> U\<close> not_in_path_image_join by blast
+            show "pathstart (pp +++ \<gamma>) = z"
+              by (simp add: \<open>pathstart pp = z\<close>)
+            show "pathfinish (pp +++ \<gamma>) = y'"
+              by (simp add: \<open>pathfinish \<gamma> = y'\<close>)
+            have paqq: "pathfinish qq = pathstart (p' \<circ> f \<circ> \<gamma>)"
+              apply (simp add: \<open>pathstart \<gamma> = y\<close> pathstart_compose)
+              apply (metis (mono_tags, lifting) \<open>l y \<in> W'\<close> \<open>path pp\<close> \<open>path qq\<close> \<open>path_image pp \<subseteq> U\<close> \<open>path_image qq \<subseteq> C\<close>
+                           \<open>pathfinish pp = y\<close> \<open>pathstart pp = z\<close> \<open>pathstart qq = a\<close>
+                           homeomorphism_apply1 [OF homUW'] l pleq pqqeq \<open>y \<in> U\<close>)
+              done
+            have "continuous_on (path_image \<gamma>) (p' \<circ> f)"
+            proof (rule continuous_on_compose)
+              show "continuous_on (path_image \<gamma>) f"
+                using \<open>path_image \<gamma> \<subseteq> V\<close> \<open>V \<subseteq> U\<close> contf continuous_on_subset by blast
+              show "continuous_on (f ` path_image \<gamma>) p'"
+                apply (rule continuous_on_subset [OF contp'])
+                apply (auto simp: path_image_def pathfinish_def pathstart_def finW)
+                done
+            qed
+            then show "path (qq +++ (p' \<circ> f \<circ> \<gamma>))"
+              using \<open>path \<gamma>\<close> \<open>path qq\<close> paqq path_continuous_image path_join_imp by blast
+            show "path_image (qq +++ (p' \<circ> f \<circ> \<gamma>)) \<subseteq> C"
+              apply (rule subset_path_image_join)
+               apply (simp add: \<open>path_image qq \<subseteq> C\<close>)
+              by (metis \<open>W' \<subseteq> C\<close> \<open>path_image \<gamma> \<subseteq> V\<close> dual_order.trans fimW(1) image_comp image_mono p'im path_image_compose)
+            show "pathstart (qq +++ (p' \<circ> f \<circ> \<gamma>)) = a"
+              by (simp add: \<open>pathstart qq = a\<close>)
+            show "p ((qq +++ (p' \<circ> f \<circ> \<gamma>)) \<xi>) = f ((pp +++ \<gamma>) \<xi>)" if \<xi>: "\<xi> \<in> {0..1}" for \<xi>
+            proof (simp add: joinpaths_def, safe)
+              show "p (qq (2*\<xi>)) = f (pp (2*\<xi>))" if "\<xi>*2 \<le> 1"
+                using \<open>\<xi> \<in> {0..1}\<close> pqqeq that by auto
+              show "p (p' (f (\<gamma> (2*\<xi> - 1)))) = f (\<gamma> (2*\<xi> - 1))" if "\<not> \<xi>*2 \<le> 1"
+                apply (rule homeomorphism_apply2 [OF homUW' finW])
+                using that \<xi> by auto
+            qed
+          qed
+          with \<open>pathfinish \<gamma> = y'\<close>  \<open>p' (f y') \<in> X\<close> show "l y' \<in> X"
+            unfolding pathfinish_join by (simp add: pathfinish_def)
+        qed
+      qed
+    qed
+    then show "continuous_on U l"
+      using openin_subopen continuous_on_open_gen [OF LC]
+      by (metis (no_types, lifting) mem_Collect_eq)
+  qed
+qed
+
+
+corollary covering_space_lift_stronger:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and f :: "'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
+      and U: "path_connected U" "locally path_connected U"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
+      and feq: "f z = p a"
+      and hom: "\<And>r. \<lbrakk>path r; path_image r \<subseteq> U; pathstart r = z; pathfinish r = z\<rbrakk>
+                     \<Longrightarrow> \<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)"
+  obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+proof (rule covering_space_lift_general [OF cov U contf fim feq])
+  fix r
+  assume "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z"
+  then obtain b where b: "homotopic_paths S (f \<circ> r) (linepath b b)"
+    using hom by blast
+  then have "f (pathstart r) = b"
+    by (metis homotopic_paths_imp_pathstart pathstart_compose pathstart_linepath)
+  then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))"
+    by (simp add: b \<open>pathstart r = z\<close>)
+  then have "homotopic_paths S (f \<circ> r) (p \<circ> linepath a a)"
+    by (simp add: o_def feq linepath_def)
+  then show "\<exists>q. path q \<and>
+                  path_image q \<subseteq> C \<and>
+                  pathstart q = a \<and> pathfinish q = a \<and> homotopic_paths S (f \<circ> r) (p \<circ> q)"
+    by (force simp: \<open>a \<in> C\<close>)
+qed auto
+
+corollary covering_space_lift_strong:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and f :: "'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S" "a \<in> C" "z \<in> U"
+      and scU: "simply_connected U" and lpcU: "locally path_connected U"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
+      and feq: "f z = p a"
+  obtains g where "continuous_on U g" "g ` U \<subseteq> C" "g z = a" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+proof (rule covering_space_lift_stronger [OF cov _ lpcU contf fim feq])
+  show "path_connected U"
+    using scU simply_connected_eq_contractible_loop_some by blast
+  fix r
+  assume r: "path r" "path_image r \<subseteq> U" "pathstart r = z" "pathfinish r = z"
+  have "linepath (f z) (f z) = f \<circ> linepath z z"
+    by (simp add: o_def linepath_def)
+  then have "homotopic_paths S (f \<circ> r) (linepath (f z) (f z))"
+    by (metis r contf fim homotopic_paths_continuous_image scU simply_connected_eq_contractible_path)
+  then show "\<exists>b. homotopic_paths S (f \<circ> r) (linepath b b)"
+    by blast
+qed blast
+
+corollary covering_space_lift:
+  fixes p :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector"
+    and f :: "'c::real_normed_vector \<Rightarrow> 'b"
+  assumes cov: "covering_space C p S"
+      and U: "simply_connected U"  "locally path_connected U"
+      and contf: "continuous_on U f" and fim: "f ` U \<subseteq> S"
+    obtains g where "continuous_on U g" "g ` U \<subseteq> C" "\<And>y. y \<in> U \<Longrightarrow> p(g y) = f y"
+proof (cases "U = {}")
+  case True
+  with that show ?thesis by auto
+next
+  case False
+  then obtain z where "z \<in> U" by blast
+  then obtain a where "a \<in> C" "f z = p a"
+    by (metis cov covering_space_imp_surjective fim image_iff image_subset_iff)
+  then show ?thesis
+    by (metis that covering_space_lift_strong [OF cov _ \<open>z \<in> U\<close> U contf fim])
+qed
+
 end
--- a/src/HOL/Analysis/Path_Connected.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Path_Connected.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -134,12 +134,18 @@
 
 subsection\<open>Basic lemmas about paths\<close>
 
+lemma continuous_on_path: "path f \<Longrightarrow> t \<subseteq> {0..1} \<Longrightarrow> continuous_on t f"
+  using continuous_on_subset path_def by blast
+
 lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g"
   by (simp add: arc_def inj_on_def simple_path_def)
 
 lemma arc_imp_path: "arc g \<Longrightarrow> path g"
   using arc_def by blast
 
+lemma arc_imp_inj_on: "arc g \<Longrightarrow> inj_on g {0..1}"
+  by (auto simp: arc_def)
+
 lemma simple_path_imp_path: "simple_path g \<Longrightarrow> path g"
   using simple_path_def by blast
 
@@ -1047,7 +1053,7 @@
     done
 qed
 
-subsection \<open>Reparametrizing a closed curve to start at some chosen point\<close>
+subsection \<open>shiftpath: Reparametrizing a closed curve to start at some chosen point\<close>
 
 definition shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
   where "shiftpath a f = (\<lambda>x. if (a + x) \<le> 1 then f (a + x) else f (a + x - 1))"
@@ -1146,6 +1152,22 @@
     by (auto simp add: image_iff)
 qed
 
+lemma simple_path_shiftpath:
+  assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \<le> a" "a \<le> 1"
+    shows "simple_path (shiftpath a g)"
+  unfolding simple_path_def
+proof (intro conjI impI ballI)
+  show "path (shiftpath a g)"
+    by (simp add: assms path_shiftpath simple_path_imp_path)
+  have *: "\<And>x y. \<lbrakk>g x = g y; x \<in> {0..1}; y \<in> {0..1}\<rbrakk> \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
+    using assms by (simp add:  simple_path_def)
+  show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
+    if "x \<in> {0..1}" "y \<in> {0..1}" "shiftpath a g x = shiftpath a g y" for x y
+    using that a unfolding shiftpath_def
+    apply (simp add: split: if_split_asm)
+      apply (drule *; auto)+
+    done
+qed
 
 subsection \<open>Special case of straight-line paths\<close>
 
@@ -1499,8 +1521,8 @@
   using path_connected_component_set by auto
 
 lemma path_connected_imp_connected:
-  assumes "path_connected s"
-  shows "connected s"
+  assumes "path_connected S"
+  shows "connected S"
   unfolding connected_def not_ex
   apply rule
   apply rule
@@ -1509,10 +1531,10 @@
   apply (elim conjE)
 proof -
   fix e1 e2
-  assume as: "open e1" "open e2" "s \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> s = {}" "e1 \<inter> s \<noteq> {}" "e2 \<inter> s \<noteq> {}"
-  then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> s" "x2 \<in> e2 \<inter> s"
+  assume as: "open e1" "open e2" "S \<subseteq> e1 \<union> e2" "e1 \<inter> e2 \<inter> S = {}" "e1 \<inter> S \<noteq> {}" "e2 \<inter> S \<noteq> {}"
+  then obtain x1 x2 where obt:"x1 \<in> e1 \<inter> S" "x2 \<in> e2 \<inter> S"
     by auto
-  then obtain g where g: "path g" "path_image g \<subseteq> s" "pathstart g = x1" "pathfinish g = x2"
+  then obtain g where g: "path g" "path_image g \<subseteq> S" "pathstart g = x1" "pathfinish g = x2"
     using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto
   have *: "connected {0..1::real}"
     by (auto intro!: convex_connected convex_real_interval)
@@ -1535,23 +1557,23 @@
 qed
 
 lemma open_path_component:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"
-  shows "open (path_component_set s x)"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "open S"
+  shows "open (path_component_set S x)"
   unfolding open_contains_ball
 proof
   fix y
-  assume as: "y \<in> path_component_set s x"
-  then have "y \<in> s"
+  assume as: "y \<in> path_component_set S x"
+  then have "y \<in> S"
     apply -
     apply (rule path_component_mem(2))
     unfolding mem_Collect_eq
     apply auto
     done
-  then obtain e where e: "e > 0" "ball y e \<subseteq> s"
+  then obtain e where e: "e > 0" "ball y e \<subseteq> S"
     using assms[unfolded open_contains_ball]
     by auto
-  show "\<exists>e > 0. ball y e \<subseteq> path_component_set s x"
+  show "\<exists>e > 0. ball y e \<subseteq> path_component_set S x"
     apply (rule_tac x=e in exI)
     apply (rule,rule \<open>e>0\<close>)
     apply rule
@@ -1559,7 +1581,7 @@
   proof -
     fix z
     assume "dist y z < e"
-    then show "path_component s x z"
+    then show "path_component S x z"
       apply (rule_tac path_component_trans[of _ _ y])
       defer
       apply (rule path_component_of_subset[OF e(2)])
@@ -1571,17 +1593,17 @@
 qed
 
 lemma open_non_path_component:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"
-  shows "open (s - path_component_set s x)"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "open S"
+  shows "open (S - path_component_set S x)"
   unfolding open_contains_ball
 proof
   fix y
-  assume as: "y \<in> s - path_component_set s x"
-  then obtain e where e: "e > 0" "ball y e \<subseteq> s"
+  assume as: "y \<in> S - path_component_set S x"
+  then obtain e where e: "e > 0" "ball y e \<subseteq> S"
     using assms [unfolded open_contains_ball]
     by auto
-  show "\<exists>e>0. ball y e \<subseteq> s - path_component_set s x"
+  show "\<exists>e>0. ball y e \<subseteq> S - path_component_set S x"
     apply (rule_tac x=e in exI)
     apply rule
     apply (rule \<open>e>0\<close>)
@@ -1590,8 +1612,8 @@
     defer
   proof (rule ccontr)
     fix z
-    assume "z \<in> ball y e" "\<not> z \<notin> path_component_set s x"
-    then have "y \<in> path_component_set s x"
+    assume "z \<in> ball y e" "\<not> z \<notin> path_component_set S x"
+    then have "y \<in> path_component_set S x"
       unfolding not_not mem_Collect_eq using \<open>e>0\<close>
       apply -
       apply (rule path_component_trans, assumption)
@@ -1605,42 +1627,42 @@
 qed
 
 lemma connected_open_path_connected:
-  fixes s :: "'a::real_normed_vector set"
-  assumes "open s"
-    and "connected s"
-  shows "path_connected s"
+  fixes S :: "'a::real_normed_vector set"
+  assumes "open S"
+    and "connected S"
+  shows "path_connected S"
   unfolding path_connected_component_set
 proof (rule, rule, rule path_component_subset, rule)
   fix x y
-  assume "x \<in> s" and "y \<in> s"
-  show "y \<in> path_component_set s x"
+  assume "x \<in> S" and "y \<in> S"
+  show "y \<in> path_component_set S x"
   proof (rule ccontr)
     assume "\<not> ?thesis"
-    moreover have "path_component_set s x \<inter> s \<noteq> {}"
-      using \<open>x \<in> s\<close> path_component_eq_empty path_component_subset[of s x]
+    moreover have "path_component_set S x \<inter> S \<noteq> {}"
+      using \<open>x \<in> S\<close> path_component_eq_empty path_component_subset[of S x]
       by auto
     ultimately
     show False
-      using \<open>y \<in> s\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
+      using \<open>y \<in> S\<close> open_non_path_component[OF assms(1)] open_path_component[OF assms(1)]
       using assms(2)[unfolded connected_def not_ex, rule_format,
-        of "path_component_set s x" "s - path_component_set s x"]
+        of "path_component_set S x" "S - path_component_set S x"]
       by auto
   qed
 qed
 
 lemma path_connected_continuous_image:
-  assumes "continuous_on s f"
-    and "path_connected s"
-  shows "path_connected (f ` s)"
+  assumes "continuous_on S f"
+    and "path_connected S"
+  shows "path_connected (f ` S)"
   unfolding path_connected_def
 proof (rule, rule)
   fix x' y'
-  assume "x' \<in> f ` s" "y' \<in> f ` s"
-  then obtain x y where x: "x \<in> s" and y: "y \<in> s" and x': "x' = f x" and y': "y' = f y"
+  assume "x' \<in> f ` S" "y' \<in> f ` S"
+  then obtain x y where x: "x \<in> S" and y: "y \<in> S" and x': "x' = f x" and y': "y' = f y"
     by auto
-  from x y obtain g where "path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y"
+  from x y obtain g where "path g \<and> path_image g \<subseteq> S \<and> pathstart g = x \<and> pathfinish g = y"
     using assms(2)[unfolded path_connected_def] by fast
-  then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` s \<and> pathstart g = x' \<and> pathfinish g = y'"
+  then show "\<exists>g. path g \<and> path_image g \<subseteq> f ` S \<and> pathstart g = x' \<and> pathfinish g = y'"
     unfolding x' y'
     apply (rule_tac x="f \<circ> g" in exI)
     unfolding path_defs
@@ -1649,12 +1671,27 @@
     done
 qed
 
-lemma path_connected_segment:
+lemma path_connected_translationI:
+  fixes a :: "'a :: topological_group_add"
+  assumes "path_connected S" shows "path_connected ((\<lambda>x. a + x) ` S)"
+  by (intro path_connected_continuous_image assms continuous_intros)
+
+lemma path_connected_translation:
+  fixes a :: "'a :: topological_group_add"
+  shows "path_connected ((\<lambda>x. a + x) ` S) = path_connected S"
+proof -
+  have "\<forall>x y. op + (x::'a) ` op + (0 - x) ` y = y"
+    by (simp add: image_image)
+  then show ?thesis
+    by (metis (no_types) path_connected_translationI)
+qed
+
+lemma path_connected_segment [simp]:
     fixes a :: "'a::real_normed_vector"
     shows "path_connected (closed_segment a b)"
   by (simp add: convex_imp_path_connected)
 
-lemma path_connected_open_segment:
+lemma path_connected_open_segment [simp]:
     fixes a :: "'a::real_normed_vector"
     shows "path_connected (open_segment a b)"
   by (simp add: convex_imp_path_connected)
@@ -1663,10 +1700,10 @@
   "s homeomorphic t \<Longrightarrow> path_connected s \<longleftrightarrow> path_connected t"
   unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image)
 
-lemma path_connected_empty: "path_connected {}"
+lemma path_connected_empty [simp]: "path_connected {}"
   unfolding path_connected_def by auto
 
-lemma path_connected_singleton: "path_connected {a}"
+lemma path_connected_singleton [simp]: "path_connected {a}"
   unfolding path_connected_def pathstart_def pathfinish_def path_image_def
   apply clarify
   apply (rule_tac x="\<lambda>x. a" in exI)
@@ -1743,7 +1780,7 @@
   unfolding path_connected_component
   by (meson path_component_path_image_pathstart path_component_sym path_component_trans)
 
-lemma path_connected_path_component:
+lemma path_connected_path_component [simp]:
    "path_connected (path_component_set s x)"
 proof -
   { fix y z
@@ -1973,54 +2010,46 @@
   by (simp add: path_connected_punctured_universe path_connected_imp_connected)
 
 lemma path_connected_sphere:
-  assumes "2 \<le> DIM('a::euclidean_space)"
-  shows "path_connected {x::'a. norm (x - a) = r}"
-proof (rule linorder_cases [of r 0])
-  assume "r < 0"
-  then have "{x::'a. norm(x - a) = r} = {}"
-    by auto
+  fixes a :: "'a :: euclidean_space"
+  assumes "2 \<le> DIM('a)"
+  shows "path_connected(sphere a r)"
+proof (cases r "0::real" rule: linorder_cases)
+  case less
   then show ?thesis
-    using path_connected_empty by simp
+    by (simp add: path_connected_empty)
 next
-  assume "r = 0"
+  case equal
   then show ?thesis
-    using path_connected_singleton by simp
+    by (simp add: path_connected_singleton)
 next
-  assume r: "0 < r"
-  have *: "{x::'a. norm(x - a) = r} = (\<lambda>x. a + r *\<^sub>R x) ` {x. norm x = 1}"
-    apply (rule set_eqI)
-    apply rule
-    unfolding image_iff
-    apply (rule_tac x="(1/r) *\<^sub>R (x - a)" in bexI)
-    unfolding mem_Collect_eq norm_scaleR
-    using r
-    apply (auto simp add: scaleR_right_diff_distrib)
-    done
-  have **: "{x::'a. norm x = 1} = (\<lambda>x. (1/norm x) *\<^sub>R x) ` (- {0})"
-    apply (rule set_eqI)
-    apply rule
-    unfolding image_iff
-    apply (rule_tac x=x in bexI)
-    unfolding mem_Collect_eq
-    apply (auto split: if_split_asm)
-    done
-  have "continuous_on (- {0}) (\<lambda>x::'a. 1 / norm x)"
-    by (auto intro!: continuous_intros)
+  case greater
+  then have eq: "(sphere (0::'a) r) = (\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a})"
+    by (force simp: image_iff split: if_split_asm)
+  have "continuous_on (- {0::'a}) (\<lambda>x. (r / norm x) *\<^sub>R x)"
+    by (intro continuous_intros) auto
+  then have "path_connected ((\<lambda>x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))"
+    by (intro path_connected_continuous_image path_connected_punctured_universe assms)
+  with eq have "path_connected (sphere (0::'a) r)"
+    by auto
+  then have "path_connected(op +a ` (sphere (0::'a) r))"
+    by (simp add: path_connected_translation)
   then show ?thesis
-    unfolding * **
-    using path_connected_punctured_universe[OF assms]
-    by (auto intro!: path_connected_continuous_image continuous_intros)
-qed
-
-corollary connected_sphere: "2 \<le> DIM('a::euclidean_space) \<Longrightarrow> connected {x::'a. norm (x - a) = r}"
-  using path_connected_sphere path_connected_imp_connected
-  by auto
+    by (metis add.right_neutral sphere_translation)
+qed
+
+lemma connected_sphere:
+    fixes a :: "'a :: euclidean_space"
+    assumes "2 \<le> DIM('a)"
+      shows "connected(sphere a r)"
+  using path_connected_sphere [OF assms]
+  by (simp add: path_connected_imp_connected)
+
 
 corollary path_connected_complement_bounded_convex:
     fixes s :: "'a :: euclidean_space set"
     assumes "bounded s" "convex s" and 2: "2 \<le> DIM('a)"
     shows "path_connected (- s)"
-proof (cases "s={}")
+proof (cases "s = {}")
   case True then show ?thesis
     using convex_imp_path_connected by auto
 next
@@ -2104,11 +2133,11 @@
     then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))"
       by (force simp: closed_segment_def intro!: path_connected_linepath)
     have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))"
-      apply (rule path_component_of_subset [of "{x. norm(x - a) = B}"])
+      apply (rule path_component_of_subset [of "sphere a B"])
        using \<open>s \<subseteq> ball a B\<close>
        apply (force simp: ball_def dist_norm norm_minus_commute)
       apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format])
-      using \<open>x \<noteq> a\<close>  using \<open>y \<noteq> a\<close>  B apply (auto simp: C_def D_def)
+       using \<open>x \<noteq> a\<close>  using \<open>y \<noteq> a\<close>  B apply (auto simp: dist_norm C_def D_def)
       done
     have "path_component (- s) x y"
       by (metis path_component_trans path_component_sym pcx pdy pyx)
@@ -2193,6 +2222,52 @@
   thus ?case by (metis Diff_insert)
 qed
 
+lemma sphere_1D_doubleton_zero:
+  assumes 1: "DIM('a) = 1" and "r > 0"
+  obtains x y::"'a::euclidean_space"
+    where "sphere 0 r = {x,y} \<and> dist x y = 2*r"
+proof -
+  obtain b::'a where b: "Basis = {b}"
+    using 1 card_1_singletonE by blast
+  show ?thesis
+  proof (intro that conjI)
+    have "x = norm x *\<^sub>R b \<or> x = - norm x *\<^sub>R b" if "r = norm x" for x
+    proof -
+      have xb: "(x \<bullet> b) *\<^sub>R b = x"
+        using euclidean_representation [of x, unfolded b] by force
+      then have "norm ((x \<bullet> b) *\<^sub>R b) = norm x"
+        by simp
+      with b have "\<bar>x \<bullet> b\<bar> = norm x"
+        using norm_Basis by fastforce
+      with xb show ?thesis
+        apply (simp add: abs_if split: if_split_asm)
+        apply (metis add.inverse_inverse real_vector.scale_minus_left xb)
+        done
+    qed
+    with \<open>r > 0\<close> b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}"
+      by (force simp: sphere_def dist_norm)
+    have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)"
+      by (simp add: dist_norm)
+    also have "... = norm ((2*r) *\<^sub>R b)"
+      by (metis mult_2 scaleR_add_left)
+    also have "... = 2*r"
+      using \<open>r > 0\<close> b norm_Basis by fastforce
+    finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" .
+  qed
+qed
+
+lemma sphere_1D_doubleton:
+  fixes a :: "'a :: euclidean_space"
+  assumes "DIM('a) = 1" and "r > 0"
+  obtains x y where "sphere a r = {x,y} \<and> dist x y = 2*r"
+proof -
+  have "sphere a r = op +a ` sphere 0 r"
+    by (metis add.right_neutral sphere_translation)
+  then show ?thesis
+    using sphere_1D_doubleton_zero [OF assms]
+    by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that)
+qed
+
 lemma psubset_sphere_Compl_connected:
   fixes S :: "'a::euclidean_space set"
   assumes S: "S \<subset> sphere a r" and "0 < r" and 2: "2 \<le> DIM('a)"
@@ -2219,6 +2294,7 @@
     by (simp add: CS connected_Un)
 qed
 
+
 subsection\<open>Relations between components and path components\<close>
 
 lemma open_connected_component:
@@ -2590,20 +2666,6 @@
 by (metis bounded_subset [of "connected_component_set (- s) _"] connected_component_maximal
        Compl_iff Un_iff assms subsetI)
 
-lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)"
-  by (simp add: Int_commute frontier_def interior_closure)
-
-lemma frontier_interior_subset: "frontier(interior S) \<subseteq> frontier S"
-  by (simp add: Diff_mono frontier_interiors interior_mono interior_subset)
-
-lemma connected_Int_frontier:
-     "\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
-  apply (simp add: frontier_interiors connected_openin, safe)
-  apply (drule_tac x="s \<inter> interior t" in spec, safe)
-   apply (drule_tac [2] x="s \<inter> interior (-t)" in spec)
-   apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD])
-  done
-
 lemma frontier_not_empty:
   fixes S :: "'a :: real_normed_vector set"
   shows "\<lbrakk>S \<noteq> {}; S \<noteq> UNIV\<rbrakk> \<Longrightarrow> frontier S \<noteq> {}"
@@ -2735,7 +2797,7 @@
     shows "connected_component_set s x = UNIV \<longleftrightarrow> s = UNIV"
   using connected_component_in connected_component_UNIV by blast
 
-lemma components_univ [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}"
+lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}"
   by (auto simp: components_eq_sing_iff)
 
 lemma interior_inside_frontier:
@@ -2837,6 +2899,26 @@
   shows "\<lbrakk>convex t; s \<subseteq> t\<rbrakk> \<Longrightarrow> - t \<subseteq> outside s"
   using outside_convex outside_mono by blast
 
+lemma outside_Un_outside_Un:
+  fixes S :: "'a::real_normed_vector set"
+  assumes "S \<inter> outside(T \<union> U) = {}"
+  shows "outside(T \<union> U) \<subseteq> outside(T \<union> S)"
+proof
+  fix x
+  assume x: "x \<in> outside (T \<union> U)"
+  have "Y \<subseteq> - S" if "connected Y" "Y \<subseteq> - T" "Y \<subseteq> - U" "x \<in> Y" "u \<in> Y" for u Y
+  proof -
+    have "Y \<subseteq> connected_component_set (- (T \<union> U)) x"
+      by (simp add: connected_component_maximal that)
+    also have "... \<subseteq> outside(T \<union> U)"
+      by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x)
+    finally have "Y \<subseteq> outside(T \<union> U)" .
+    with assms show ?thesis by auto
+  qed
+  with x show "x \<in> outside (T \<union> S)"
+    by (simp add: outside_connected_component_lt connected_component_def) meson
+qed
+
 lemma outside_frontier_misses_closure:
     fixes s :: "'a::real_normed_vector set"
     assumes "bounded s"
@@ -3560,6 +3642,56 @@
   by (metis homotopic_compose_continuous_right homotopic_with_imp_continuous homotopic_with_imp_subset2)
 
 
+text\<open>Homotopic triviality implicitly incorporates path-connectedness.\<close>
+lemma homotopic_triviality:
+  fixes S :: "'a::real_normed_vector set"
+  shows  "(\<forall>f g. continuous_on S f \<and> f ` S \<subseteq> T \<and>
+                 continuous_on S g \<and> g ` S \<subseteq> T
+                 \<longrightarrow> homotopic_with (\<lambda>x. True) S T f g) \<longleftrightarrow>
+          (S = {} \<or> path_connected T) \<and>
+          (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> T \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)))"
+          (is "?lhs = ?rhs")
+proof (cases "S = {} \<or> T = {}")
+  case True then show ?thesis by auto
+next
+  case False show ?thesis
+  proof
+    assume LHS [rule_format]: ?lhs
+    have pab: "path_component T a b" if "a \<in> T" "b \<in> T" for a b
+    proof -
+      have "homotopic_with (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. b)"
+        by (simp add: LHS continuous_on_const image_subset_iff that)
+      then show ?thesis
+        using False homotopic_constant_maps by blast
+    qed
+      moreover
+    have "\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)" if "continuous_on S f" "f ` S \<subseteq> T" for f
+      by (metis (full_types) False LHS equals0I homotopic_constant_maps homotopic_with_imp_continuous homotopic_with_imp_subset2 pab that)
+    ultimately show ?rhs
+      by (simp add: path_connected_component)
+  next
+    assume RHS: ?rhs
+    with False have T: "path_connected T"
+      by blast
+    show ?lhs
+    proof clarify
+      fix f g
+      assume "continuous_on S f" "f ` S \<subseteq> T" "continuous_on S g" "g ` S \<subseteq> T"
+      obtain c d where c: "homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)" and d: "homotopic_with (\<lambda>x. True) S T g (\<lambda>x. d)"
+        using False \<open>continuous_on S f\<close> \<open>f ` S \<subseteq> T\<close>  RHS \<open>continuous_on S g\<close> \<open>g ` S \<subseteq> T\<close> by blast
+      then have "c \<in> T" "d \<in> T"
+        using False homotopic_with_imp_subset2 by fastforce+
+      with T have "path_component T c d"
+        using path_connected_component by blast
+      then have "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) (\<lambda>x. d)"
+        by (simp add: homotopic_constant_maps)
+      with c d show "homotopic_with (\<lambda>x. True) S T f g"
+        by (meson homotopic_with_symD homotopic_with_trans)
+    qed
+  qed
+qed
+
+
 subsection\<open>Homotopy of paths, maintaining the same endpoints.\<close>
 
 
@@ -4519,7 +4651,7 @@
   shows "starlike S \<Longrightarrow> contractible S"
 using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def)
 
-lemma contractible_UNIV: "contractible (UNIV :: 'a::real_normed_vector set)"
+lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)"
   by (simp add: starlike_imp_contractible)
 
 lemma starlike_imp_simply_connected:
@@ -4547,8 +4679,8 @@
   shows "is_interval S \<longleftrightarrow> simply_connected S"
 using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto
 
-lemma contractible_empty: "contractible {}"
-  by (simp add: continuous_on_empty contractible_def homotopic_with)
+lemma contractible_empty [simp]: "contractible {}"
+  by (simp add: contractible_def homotopic_with)
 
 lemma contractible_convex_tweak_boundary_points:
   fixes S :: "'a::euclidean_space set"
@@ -4557,7 +4689,7 @@
 proof (cases "S = {}")
   case True
   with assms show ?thesis
-    by (simp add: contractible_empty subsetCE)
+    by (simp add: subsetCE)
 next
   case False
   show ?thesis
@@ -4569,9 +4701,9 @@
 lemma convex_imp_contractible:
   fixes S :: "'a::real_normed_vector set"
   shows "convex S \<Longrightarrow> contractible S"
-using contractible_empty convex_imp_starlike starlike_imp_contractible by auto
-
-lemma contractible_sing:
+  using contractible_empty convex_imp_starlike starlike_imp_contractible by blast
+
+lemma contractible_sing [simp]:
   fixes a :: "'a::real_normed_vector"
   shows "contractible {a}"
 by (rule convex_imp_contractible [OF convex_singleton])
@@ -4705,6 +4837,28 @@
 apply (auto intro: P)
 done
 
+lemma locally_Times:
+  fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set"
+  assumes PS: "locally P S" and QT: "locally Q T" and R: "\<And>S T. P S \<and> Q T \<Longrightarrow> R(S \<times> T)"
+  shows "locally R (S \<times> T)"
+    unfolding locally_def
+proof (clarify)
+  fix W x y
+  assume W: "openin (subtopology euclidean (S \<times> T)) W" and xy: "(x, y) \<in> W"
+  then obtain U V where "openin (subtopology euclidean S) U" "x \<in> U"
+                        "openin (subtopology euclidean T) V" "y \<in> V" "U \<times> V \<subseteq> W"
+    using Times_in_interior_subtopology by metis
+  then obtain U1 U2 V1 V2
+         where opeS: "openin (subtopology euclidean S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U"
+           and opeT: "openin (subtopology euclidean T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V"
+    by (meson PS QT locallyE)
+  with \<open>U \<times> V \<subseteq> W\<close> show "\<exists>u v. openin (subtopology euclidean (S \<times> T)) u \<and> R v \<and> (x,y) \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> W"
+    apply (rule_tac x="U1 \<times> V1" in exI)
+    apply (rule_tac x="U2 \<times> V2" in exI)
+    apply (auto simp: openin_Times R)
+    done
+qed
+
 
 proposition homeomorphism_locally_imp:
   fixes S :: "'a::metric_space set" and t :: "'b::t2_space set"
@@ -5148,6 +5302,106 @@
     using compact_eq_bounded_closed locally_mono by blast
 qed
 
+lemma locally_compact_openin_Un:
+  fixes S :: "'a::euclidean_space set"
+  assumes LCS: "locally compact S" and LCT:"locally compact T"
+      and opS: "openin (subtopology euclidean (S \<union> T)) S"
+      and opT: "openin (subtopology euclidean (S \<union> T)) T"
+    shows "locally compact (S \<union> T)"
+proof -
+  have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" for x
+  proof -
+    obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
+      using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
+    moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> S"
+      by (meson \<open>x \<in> S\<close> opS openin_contains_cball)
+    then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> S"
+      by force
+    ultimately show ?thesis
+      apply (rule_tac x="min e1 e2" in exI)
+      apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int)
+      by (metis closed_Int closed_cball inf_left_commute)
+  qed
+  moreover have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> T" for x
+  proof -
+    obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)"
+      using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
+    moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> T"
+      by (meson \<open>x \<in> T\<close> opT openin_contains_cball)
+    then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> T"
+      by force
+    ultimately show ?thesis
+      apply (rule_tac x="min e1 e2" in exI)
+      apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int)
+      by (metis closed_Int closed_cball inf_left_commute)
+  qed
+  ultimately show ?thesis
+    by (force simp: locally_compact_Int_cball)
+qed
+
+lemma locally_compact_closedin_Un:
+  fixes S :: "'a::euclidean_space set"
+  assumes LCS: "locally compact S" and LCT:"locally compact T"
+      and clS: "closedin (subtopology euclidean (S \<union> T)) S"
+      and clT: "closedin (subtopology euclidean (S \<union> T)) T"
+    shows "locally compact (S \<union> T)"
+proof -
+  have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x
+  proof -
+    obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
+      using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
+    moreover
+    obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)"
+      using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
+    ultimately show ?thesis
+      apply (rule_tac x="min e1 e2" in exI)
+      apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
+      by (metis closed_Int closed_Un closed_cball inf_left_commute)
+  qed
+  moreover
+  have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<in> S" "x \<notin> T" for x
+  proof -
+    obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)"
+      using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast
+    moreover
+    obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S - T"
+      using clT x by (fastforce simp: openin_contains_cball closedin_def)
+    then have "closed (cball x e2 \<inter> T)"
+    proof -
+      have "{} = T - (T - cball x e2)"
+        using Diff_subset Int_Diff \<open>cball x e2 \<inter> (S \<union> T) \<subseteq> S - T\<close> by auto
+      then show ?thesis
+        by (simp add: Diff_Diff_Int inf_commute)
+    qed
+    ultimately show ?thesis
+      apply (rule_tac x="min e1 e2" in exI)
+      apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
+      by (metis closed_Int closed_Un closed_cball inf_left_commute)
+  qed
+  moreover
+  have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<notin> S" "x \<in> T" for x
+  proof -
+    obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)"
+      using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast
+    moreover
+    obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S \<union> T - S"
+      using clS x by (fastforce simp: openin_contains_cball closedin_def)
+    then have "closed (cball x e2 \<inter> S)"
+      by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb)
+    ultimately show ?thesis
+      apply (rule_tac x="min e1 e2" in exI)
+      apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib)
+      by (metis closed_Int closed_Un closed_cball inf_left_commute)
+  qed
+  ultimately show ?thesis
+    by (auto simp: locally_compact_Int_cball)
+qed
+
+lemma locally_compact_Times:
+  fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set"
+  shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<times> T)"
+  by (auto simp: compact_Times locally_Times)
+
 subsection\<open>Important special cases of local connectedness and path connectedness\<close>
 
 lemma locally_connected_1:
@@ -6519,6 +6773,64 @@
     by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable)
 qed
 
+lemma connected_uncountable:
+  fixes S :: "'a::metric_space set"
+  assumes "connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S"
+proof -
+  have "continuous_on S (dist a)"
+    by (intro continuous_intros)
+  then have "connected (dist a ` S)"
+    by (metis connected_continuous_image \<open>connected S\<close>)
+  then have "closed_segment 0 (dist a b) \<subseteq> (dist a ` S)"
+    by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex)
+  then have "uncountable (dist a ` S)"
+    by (metis \<open>a \<noteq> b\<close> countable_subset dist_eq_0_iff uncountable_closed_segment)
+  then show ?thesis
+    by blast
+qed
+
+lemma path_connected_uncountable:
+  fixes S :: "'a::metric_space set"
+  assumes "path_connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S"
+  using path_connected_imp_connected assms connected_uncountable by metis
+
+lemma connected_finite_iff_sing:
+  fixes S :: "'a::metric_space set"
+  assumes "connected S"
+  shows "finite S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})"  (is "_ = ?rhs")
+proof -
+  have "uncountable S" if "\<not> ?rhs"
+    using connected_uncountable assms that by blast
+  then show ?thesis
+    using uncountable_infinite by auto
+qed
+
+lemma connected_card_eq_iff_nontrivial:
+  fixes S :: "'a::metric_space set"
+  shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> ~(\<exists>a. S \<subseteq> {a})"
+  apply (auto simp: countable_finite finite_subset)
+  by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff)
+
+lemma simple_path_image_uncountable:
+  fixes g :: "real \<Rightarrow> 'a::metric_space"
+  assumes "simple_path g"
+  shows "uncountable (path_image g)"
+proof -
+  have "g 0 \<in> path_image g" "g (1/2) \<in> path_image g"
+    by (simp_all add: path_defs)
+  moreover have "g 0 \<noteq> g (1/2)"
+    using assms by (fastforce simp add: simple_path_def)
+  ultimately show ?thesis
+    apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image)
+    by blast
+qed
+
+lemma arc_image_uncountable:
+  fixes g :: "real \<Rightarrow> 'a::metric_space"
+  assumes "arc g"
+  shows "uncountable (path_image g)"
+  by (simp add: arc_imp_simple_path assms simple_path_image_uncountable)
+
 
 subsection\<open> Some simple positive connection theorems\<close>
 
--- a/src/HOL/Analysis/Topology_Euclidean_Space.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -15,24 +15,6 @@
   Norm_Arith
 begin
 
-(* FIXME: move to HOL/Real_Vector_Spaces.thy *)
-
-lemma scaleR_2:
-  fixes x :: "'a::real_vector"
-  shows "scaleR 2 x = x + x"
-  unfolding one_add_one [symmetric] scaleR_left_distrib by simp
-
-lemma scaleR_half_double [simp]:
-  fixes a :: "'a::real_normed_vector"
-  shows "(1 / 2) *\<^sub>R (a + a) = a"
-proof -
-  have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a"
-    by (metis scaleR_2 scaleR_scaleR)
-  then show ?thesis
-    by simp
-qed
-
-
 (* FIXME: move elsewhere *)
 
 lemma Times_eq_image_sum:
@@ -894,6 +876,11 @@
 lemma closed_subset: "S \<subseteq> T \<Longrightarrow> closed S \<Longrightarrow> closedin (subtopology euclidean T) S"
   by (auto simp add: closedin_closed)
 
+lemma closedin_closed_subset:
+ "\<lbrakk>closedin (subtopology euclidean U) V; T \<subseteq> U; S = V \<inter> T\<rbrakk>
+             \<Longrightarrow> closedin (subtopology euclidean T) S"
+  by (metis (no_types, lifting) Int_assoc Int_commute closedin_closed inf.orderE)
+
 lemma finite_imp_closedin:
   fixes S :: "'a::t1_space set"
   shows "\<lbrakk>finite S; S \<subseteq> T\<rbrakk> \<Longrightarrow> closedin (subtopology euclidean T) S"
@@ -1099,7 +1086,7 @@
 lemma cball_min_Int: "cball a (min r s) = cball a r \<inter> cball a s"
   by (simp add: set_eq_iff)
 
-lemma cball_diff_eq_sphere: "cball a r - ball a r =  {x. dist x a = r}"
+lemma cball_diff_eq_sphere: "cball a r - ball a r =  sphere a r"
   by (auto simp: cball_def ball_def dist_commute)
 
 lemma image_add_ball [simp]:
@@ -2744,6 +2731,28 @@
 lemma frontier_UNIV [simp]: "frontier UNIV = {}"
   using frontier_complement frontier_empty by fastforce
 
+lemma frontier_interiors: "frontier s = - interior(s) - interior(-s)"
+  by (simp add: Int_commute frontier_def interior_closure)
+
+lemma frontier_interior_subset: "frontier(interior S) \<subseteq> frontier S"
+  by (simp add: Diff_mono frontier_interiors interior_mono interior_subset)
+
+lemma connected_Int_frontier:
+     "\<lbrakk>connected s; s \<inter> t \<noteq> {}; s - t \<noteq> {}\<rbrakk> \<Longrightarrow> (s \<inter> frontier t \<noteq> {})"
+  apply (simp add: frontier_interiors connected_openin, safe)
+  apply (drule_tac x="s \<inter> interior t" in spec, safe)
+   apply (drule_tac [2] x="s \<inter> interior (-t)" in spec)
+   apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD])
+  done
+
+lemma closure_Un_frontier: "closure S = S \<union> frontier S"
+proof -
+  have "S \<union> interior S = S"
+    using interior_subset by auto
+  then show ?thesis
+    using closure_subset by (auto simp: frontier_def)
+qed
+
 
 subsection \<open>Filters and the ``eventually true'' quantifier\<close>
 
--- a/src/HOL/Library/Polynomial.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Library/Polynomial.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -12,6 +12,21 @@
   "~~/src/HOL/Library/Infinite_Set"
 begin
 
+subsection \<open>Misc\<close>
+
+lemma quotient_of_denom_pos': "snd (quotient_of x) > 0"
+  using quotient_of_denom_pos [OF surjective_pairing] .
+  
+lemma of_int_divide_in_Ints: 
+  "b dvd a \<Longrightarrow> of_int a div of_int b \<in> (\<int> :: 'a :: idom_divide set)"
+proof (cases "of_int b = (0 :: 'a)")
+  case False
+  assume "b dvd a"
+  then obtain c where "a = b * c" ..
+  with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
+qed auto
+
+  
 subsection \<open>Auxiliary: operations for lists (later) representing coefficients\<close>
 
 definition cCons :: "'a::zero \<Rightarrow> 'a list \<Rightarrow> 'a list"  (infixr "##" 65)
@@ -143,6 +158,33 @@
   "coeff p (degree p) = 0 \<longleftrightarrow> p = 0"
   by (cases "p = 0", simp, simp add: leading_coeff_neq_0)
 
+lemma eq_zero_or_degree_less:
+  assumes "degree p \<le> n" and "coeff p n = 0"
+  shows "p = 0 \<or> degree p < n"
+proof (cases n)
+  case 0
+  with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close>
+  have "coeff p (degree p) = 0" by simp
+  then have "p = 0" by simp
+  then show ?thesis ..
+next
+  case (Suc m)
+  have "\<forall>i>n. coeff p i = 0"
+    using \<open>degree p \<le> n\<close> by (simp add: coeff_eq_0)
+  then have "\<forall>i\<ge>n. coeff p i = 0"
+    using \<open>coeff p n = 0\<close> by (simp add: le_less)
+  then have "\<forall>i>m. coeff p i = 0"
+    using \<open>n = Suc m\<close> by (simp add: less_eq_Suc_le)
+  then have "degree p \<le> m"
+    by (rule degree_le)
+  then have "degree p < n"
+    using \<open>n = Suc m\<close> by (simp add: less_Suc_eq_le)
+  then show ?thesis ..
+qed
+
+lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \<Longrightarrow> degree rrr \<le> dr \<Longrightarrow> degree rrr \<le> dr - 1"
+  using eq_zero_or_degree_less by fastforce
+
 
 subsection \<open>List-style constructor for polynomials\<close>
 
@@ -481,6 +523,7 @@
   "p \<noteq> 0 \<Longrightarrow> fold_coeffs f (pCons a p) = f a \<circ> fold_coeffs f p"
   by (simp add: fold_coeffs_def)
 
+
 subsection \<open>Canonical morphism on polynomials -- evaluation\<close>
 
 definition poly :: "'a::comm_semiring_0 poly \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -572,8 +615,22 @@
   
 lemma monom_eq_const_iff: "monom c n = [:d:] \<longleftrightarrow> c = d \<and> (c = 0 \<or> n = 0)"
   using monom_eq_iff'[of c n d 0] by (simp add: monom_0)
-  
-    
+
+
+subsection \<open>Leading coefficient\<close>
+
+abbreviation lead_coeff:: "'a::zero poly \<Rightarrow> 'a"
+  where "lead_coeff p \<equiv> coeff p (degree p)"
+
+lemma lead_coeff_pCons[simp]:
+  "p \<noteq> 0 \<Longrightarrow> lead_coeff (pCons a p) = lead_coeff p"
+  "p = 0 \<Longrightarrow> lead_coeff (pCons a p) = a"
+  by auto
+
+lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c"
+  by (cases "c = 0") (simp_all add: degree_monom_eq)
+
+
 subsection \<open>Addition and subtraction\<close>
 
 instantiation poly :: (comm_monoid_add) comm_monoid_add
@@ -694,6 +751,16 @@
   "degree (- p) = degree p"
   unfolding degree_def by simp
 
+lemma lead_coeff_add_le:
+  assumes "degree p < degree q"
+  shows "lead_coeff (p + q) = lead_coeff q" 
+  using assms
+  by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right)
+
+lemma lead_coeff_minus:
+  "lead_coeff (- p) = - lead_coeff p"
+  by (metis coeff_minus degree_minus)
+
 lemma degree_diff_le_max:
   fixes p q :: "'a :: ab_group_add poly"
   shows "degree (p - q) \<le> max (degree p) (degree q)"
@@ -894,7 +961,16 @@
   shows "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))"
   by (rule coeffs_eqI)
     (auto simp add: not_0_coeffs_not_Nil last_map last_coeffs_not_0 nth_default_map_eq nth_default_coeffs_eq)
-   
+
+lemma smult_eq_iff:
+  assumes "(b :: 'a :: field) \<noteq> 0"
+  shows   "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
+proof
+  assume "smult a p = smult b q"
+  also from assms have "smult (inverse b) \<dots> = q" by simp
+  finally show "smult (a / b) p = q" by (simp add: field_simps)
+qed (insert assms, auto)
+
 instantiation poly :: (comm_semiring_0) comm_semiring_0
 begin
 
@@ -1037,6 +1113,10 @@
   "degree (p ^ n) \<le> degree p * n"
   by (induct n) (auto intro: order_trans degree_mult_le)
 
+lemma coeff_0_power:
+  "coeff (p ^ n) 0 = coeff p 0 ^ n"
+  by (induction n) (simp_all add: coeff_mult)
+
 lemma poly_smult [simp]:
   "poly (smult a p) x = a * poly p x"
   by (induct p, simp, simp add: algebra_simps)
@@ -1064,6 +1144,40 @@
     by (rule le_trans[OF degree_mult_le], insert insert, auto)
 qed simp
 
+lemma coeff_0_prod_list:
+  "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
+  by (induction xs) (simp_all add: coeff_mult)
+
+lemma coeff_monom_mult: 
+  "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))"
+proof -
+  have "coeff (monom c n * p) k = (\<Sum>i\<le>k. (if n = i then c else 0) * coeff p (k - i))"
+    by (simp add: coeff_mult)
+  also have "\<dots> = (\<Sum>i\<le>k. (if n = i then c * coeff p (k - i) else 0))"
+    by (intro sum.cong) simp_all
+  also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))" by (simp add: sum.delta')
+  finally show ?thesis .
+qed
+
+lemma monom_1_dvd_iff':
+  "monom 1 n dvd p \<longleftrightarrow> (\<forall>k<n. coeff p k = 0)"
+proof
+  assume "monom 1 n dvd p"
+  then obtain r where r: "p = monom 1 n * r" by (elim dvdE)
+  thus "\<forall>k<n. coeff p k = 0" by (simp add: coeff_mult)
+next
+  assume zero: "(\<forall>k<n. coeff p k = 0)"
+  define r where "r = Abs_poly (\<lambda>k. coeff p (k + n))"
+  have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0"
+    by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, 
+        subst cofinite_eq_sequentially [symmetric]) transfer
+  hence coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def
+    by (subst poly.Abs_poly_inverse) simp_all
+  have "p = monom 1 n * r"
+    by (intro poly_eqI, subst coeff_monom_mult) (insert zero, simp_all)
+  thus "monom 1 n dvd p" by simp
+qed
+
 
 subsection \<open>Mapping polynomials\<close>
 
@@ -1169,7 +1283,7 @@
   by (intro poly_eqI) (simp_all add: coeff_map_poly)
 
 
-subsection \<open>Conversions from natural numbers\<close>
+subsection \<open>Conversions from @{typ nat} and @{typ int} numbers\<close>
 
 lemma of_nat_poly: "of_nat n = [:of_nat n :: 'a :: comm_semiring_1:]"
 proof (induction n)
@@ -1185,11 +1299,29 @@
 lemma degree_of_nat [simp]: "degree (of_nat n) = 0"
   by (simp add: of_nat_poly)
 
-lemma degree_numeral [simp]: "degree (numeral n) = 0"
-  by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
+lemma lead_coeff_of_nat [simp]:
+  "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})"
+  by (simp add: of_nat_poly)
+
+lemma of_int_poly: "of_int k = [:of_int k :: 'a :: comm_ring_1:]"
+  by (simp only: of_int_of_nat of_nat_poly) simp
+
+lemma degree_of_int [simp]: "degree (of_int k) = 0"
+  by (simp add: of_int_poly)
+
+lemma lead_coeff_of_int [simp]:
+  "lead_coeff (of_int k) = (of_int k :: 'a :: {comm_ring_1,ring_char_0})"
+  by (simp add: of_int_poly)
 
 lemma numeral_poly: "numeral n = [:numeral n:]"
   by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
+    
+lemma degree_numeral [simp]: "degree (numeral n) = 0"
+  by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
+
+lemma lead_coeff_numeral [simp]: 
+  "lead_coeff (numeral n) = numeral n"
+  by (simp add: numeral_poly)
 
 
 subsection \<open>Lemmas about divisibility\<close>
@@ -1231,6 +1363,28 @@
   shows "smult a p dvd q \<longleftrightarrow> (if a = 0 then q = 0 else p dvd q)"
   by (auto elim: smult_dvd smult_dvd_cancel)
 
+lemma is_unit_smult_iff: "smult c p dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
+proof -
+  have "smult c p = [:c:] * p" by simp
+  also have "\<dots> dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
+  proof safe
+    assume A: "[:c:] * p dvd 1"
+    thus "p dvd 1" by (rule dvd_mult_right)
+    from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE)
+    have "c dvd c * (coeff p 0 * coeff q 0)" by simp
+    also have "\<dots> = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult)
+    also note B [symmetric]
+    finally show "c dvd 1" by simp
+  next
+    assume "c dvd 1" "p dvd 1"
+    from \<open>c dvd 1\<close> obtain d where "1 = c * d" by (erule dvdE)
+    hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac)
+    hence "[:c:] dvd 1" by (rule dvdI)
+    from mult_dvd_mono[OF this \<open>p dvd 1\<close>] show "[:c:] * p dvd 1" by simp
+  qed
+  finally show ?thesis .
+qed
+
 
 subsection \<open>Polynomials form an integral domain\<close>
 
@@ -1296,6 +1450,27 @@
   "[:a::'a::{comm_semiring_1,semiring_no_zero_divisors}:] dvd [:b:] \<longleftrightarrow> a dvd b"
   by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits)
 
+lemma lead_coeff_mult:
+  fixes p q :: "'a :: {comm_semiring_0, semiring_no_zero_divisors} poly"
+  shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
+  by (cases "p=0 \<or> q=0", auto simp add:coeff_mult_degree_sum degree_mult_eq)
+
+lemma lead_coeff_smult:
+  "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p"
+proof -
+  have "smult c p = [:c:] * p" by simp
+  also have "lead_coeff \<dots> = c * lead_coeff p"
+    by (subst lead_coeff_mult) simp_all
+  finally show ?thesis .
+qed
+
+lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1"
+  by simp
+
+lemma lead_coeff_power: 
+  "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n"
+  by (induction n) (simp_all add: lead_coeff_mult)
+
 
 subsection \<open>Polynomials form an ordered integral domain\<close>
 
@@ -1401,77 +1576,10 @@
 text \<open>TODO: Simplification rules for comparisons\<close>
 
 
-subsection \<open>Leading coefficient\<close>
-
-definition lead_coeff:: "'a::zero poly \<Rightarrow> 'a" where
-  "lead_coeff p= coeff p (degree p)"
-
-lemma lead_coeff_pCons[simp]:
-    "p\<noteq>0 \<Longrightarrow>lead_coeff (pCons a p) = lead_coeff p"
-    "p=0 \<Longrightarrow> lead_coeff (pCons a p) = a"
-unfolding lead_coeff_def by auto
-
-lemma lead_coeff_0[simp]:"lead_coeff 0 =0" 
-  unfolding lead_coeff_def by auto
-
-lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\<lambda>p. coeff p 0) xs)"
-  by (induction xs) (simp_all add: coeff_mult)
-
-lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n"
-  by (induction n) (simp_all add: coeff_mult)
-
-lemma lead_coeff_mult:
-   fixes p q::"'a :: {comm_semiring_0,semiring_no_zero_divisors} poly"
-   shows "lead_coeff (p * q) = lead_coeff p * lead_coeff q"
-by (unfold lead_coeff_def,cases "p=0 \<or> q=0",auto simp add:coeff_mult_degree_sum degree_mult_eq)
-
-lemma lead_coeff_add_le:
-  assumes "degree p < degree q"
-  shows "lead_coeff (p+q) = lead_coeff q" 
-using assms unfolding lead_coeff_def
-by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right)
-
-lemma lead_coeff_minus:
-  "lead_coeff (-p) = - lead_coeff p"
-by (metis coeff_minus degree_minus lead_coeff_def)
-
-lemma lead_coeff_smult:
-  "lead_coeff (smult c p :: 'a :: {comm_semiring_0,semiring_no_zero_divisors} poly) = c * lead_coeff p"
-proof -
-  have "smult c p = [:c:] * p" by simp
-  also have "lead_coeff \<dots> = c * lead_coeff p"
-    by (subst lead_coeff_mult) simp_all
-  finally show ?thesis .
-qed
-
-lemma lead_coeff_eq_zero_iff [simp]: "lead_coeff p = 0 \<longleftrightarrow> p = 0"
-  by (simp add: lead_coeff_def)
-
-lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1"
-  by (simp add: lead_coeff_def)
-
-lemma lead_coeff_of_nat [simp]:
-  "lead_coeff (of_nat n) = (of_nat n :: 'a :: {comm_semiring_1,semiring_char_0})"
-  by (induction n) (simp_all add: lead_coeff_def of_nat_poly)
-
-lemma lead_coeff_numeral [simp]: 
-  "lead_coeff (numeral n) = numeral n"
-  unfolding lead_coeff_def
-  by (subst of_nat_numeral [symmetric], subst of_nat_poly) simp
-
-lemma lead_coeff_power: 
-  "lead_coeff (p ^ n :: 'a :: {comm_semiring_1,semiring_no_zero_divisors} poly) = lead_coeff p ^ n"
-  by (induction n) (simp_all add: lead_coeff_mult)
-
-lemma lead_coeff_nonzero: "p \<noteq> 0 \<Longrightarrow> lead_coeff p \<noteq> 0"
-  by (simp add: lead_coeff_def)
-
-lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c"
-  by (cases "c = 0") (simp_all add: lead_coeff_def degree_monom_eq)
-
-
 subsection \<open>Synthetic division and polynomial roots\<close>
 
+subsubsection \<open>Synthetic division\<close>  
+  
 text \<open>
   Synthetic division is simply division by the linear polynomial @{term "x - c"}.
 \<close>
@@ -1539,9 +1647,12 @@
   using synthetic_div_correct [of p c]
   by (simp add: algebra_simps)
 
+    
+subsubsection \<open>Polynomial roots\<close>
+  
 lemma poly_eq_0_iff_dvd:
   fixes c :: "'a::{comm_ring_1}"
-  shows "poly p c = 0 \<longleftrightarrow> [:-c, 1:] dvd p"
+  shows "poly p c = 0 \<longleftrightarrow> [:- c, 1:] dvd p"
 proof
   assume "poly p c = 0"
   with synthetic_div_correct' [of c p]
@@ -1555,7 +1666,7 @@
 
 lemma dvd_iff_poly_eq_0:
   fixes c :: "'a::{comm_ring_1}"
-  shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (-c) = 0"
+  shows "[:c, 1:] dvd p \<longleftrightarrow> poly p (- c) = 0"
   by (simp add: poly_eq_0_iff_dvd)
 
 lemma poly_roots_finite:
@@ -1610,1318 +1721,8 @@
   shows "(\<forall>x. poly p x = 0) \<longleftrightarrow> p = 0"
   by (auto simp add: poly_eq_poly_eq_iff [symmetric])
 
-
-subsection \<open>Long division of polynomials\<close>
-
-inductive eucl_rel_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly \<Rightarrow> bool"
-  where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)"
-  | eucl_rel_poly_dividesI: "y \<noteq> 0 \<Longrightarrow> x = q * y \<Longrightarrow> eucl_rel_poly x y (q, 0)"
-  | eucl_rel_poly_remainderI: "y \<noteq> 0 \<Longrightarrow> degree r < degree y
-      \<Longrightarrow> x = q * y + r \<Longrightarrow> eucl_rel_poly x y (q, r)"
-  
-lemma eucl_rel_poly_iff:
-  "eucl_rel_poly x y (q, r) \<longleftrightarrow>
-    x = q * y + r \<and>
-      (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)"
-  by (auto elim: eucl_rel_poly.cases
-    intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI)
-  
-lemma eucl_rel_poly_0:
-  "eucl_rel_poly 0 y (0, 0)"
-  unfolding eucl_rel_poly_iff by simp
-
-lemma eucl_rel_poly_by_0:
-  "eucl_rel_poly x 0 (0, x)"
-  unfolding eucl_rel_poly_iff by simp
-
-lemma eq_zero_or_degree_less:
-  assumes "degree p \<le> n" and "coeff p n = 0"
-  shows "p = 0 \<or> degree p < n"
-proof (cases n)
-  case 0
-  with \<open>degree p \<le> n\<close> and \<open>coeff p n = 0\<close>
-  have "coeff p (degree p) = 0" by simp
-  then have "p = 0" by simp
-  then show ?thesis ..
-next
-  case (Suc m)
-  have "\<forall>i>n. coeff p i = 0"
-    using \<open>degree p \<le> n\<close> by (simp add: coeff_eq_0)
-  then have "\<forall>i\<ge>n. coeff p i = 0"
-    using \<open>coeff p n = 0\<close> by (simp add: le_less)
-  then have "\<forall>i>m. coeff p i = 0"
-    using \<open>n = Suc m\<close> by (simp add: less_eq_Suc_le)
-  then have "degree p \<le> m"
-    by (rule degree_le)
-  then have "degree p < n"
-    using \<open>n = Suc m\<close> by (simp add: less_Suc_eq_le)
-  then show ?thesis ..
-qed
-
-lemma eucl_rel_poly_pCons:
-  assumes rel: "eucl_rel_poly x y (q, r)"
-  assumes y: "y \<noteq> 0"
-  assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)"
-  shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)"
-    (is "eucl_rel_poly ?x y (?q, ?r)")
-proof -
-  have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
-    using assms unfolding eucl_rel_poly_iff by simp_all
-
-  have 1: "?x = ?q * y + ?r"
-    using b x by simp
-
-  have 2: "?r = 0 \<or> degree ?r < degree y"
-  proof (rule eq_zero_or_degree_less)
-    show "degree ?r \<le> degree y"
-    proof (rule degree_diff_le)
-      show "degree (pCons a r) \<le> degree y"
-        using r by auto
-      show "degree (smult b y) \<le> degree y"
-        by (rule degree_smult_le)
-    qed
-  next
-    show "coeff ?r (degree y) = 0"
-      using \<open>y \<noteq> 0\<close> unfolding b by simp
-  qed
-
-  from 1 2 show ?thesis
-    unfolding eucl_rel_poly_iff
-    using \<open>y \<noteq> 0\<close> by simp
-qed
-
-lemma eucl_rel_poly_exists: "\<exists>q r. eucl_rel_poly x y (q, r)"
-apply (cases "y = 0")
-apply (fast intro!: eucl_rel_poly_by_0)
-apply (induct x)
-apply (fast intro!: eucl_rel_poly_0)
-apply (fast intro!: eucl_rel_poly_pCons)
-done
-
-lemma eucl_rel_poly_unique:
-  assumes 1: "eucl_rel_poly x y (q1, r1)"
-  assumes 2: "eucl_rel_poly x y (q2, r2)"
-  shows "q1 = q2 \<and> r1 = r2"
-proof (cases "y = 0")
-  assume "y = 0" with assms show ?thesis
-    by (simp add: eucl_rel_poly_iff)
-next
-  assume [simp]: "y \<noteq> 0"
-  from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y"
-    unfolding eucl_rel_poly_iff by simp_all
-  from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y"
-    unfolding eucl_rel_poly_iff by simp_all
-  from q1 q2 have q3: "(q1 - q2) * y = r2 - r1"
-    by (simp add: algebra_simps)
-  from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y"
-    by (auto intro: degree_diff_less)
-
-  show "q1 = q2 \<and> r1 = r2"
-  proof (rule ccontr)
-    assume "\<not> (q1 = q2 \<and> r1 = r2)"
-    with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto
-    with r3 have "degree (r2 - r1) < degree y" by simp
-    also have "degree y \<le> degree (q1 - q2) + degree y" by simp
-    also have "\<dots> = degree ((q1 - q2) * y)"
-      using \<open>q1 \<noteq> q2\<close> by (simp add: degree_mult_eq)
-    also have "\<dots> = degree (r2 - r1)"
-      using q3 by simp
-    finally have "degree (r2 - r1) < degree (r2 - r1)" .
-    then show "False" by simp
-  qed
-qed
-
-lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \<longleftrightarrow> q = 0 \<and> r = 0"
-by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0)
-
-lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \<longleftrightarrow> q = 0 \<and> r = x"
-by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0)
-
-lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1]
-
-lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2]
-
-
-
-subsection \<open>Pseudo-Division and Division of Polynomials\<close>
-
-text\<open>This part is by René Thiemann and Akihisa Yamada.\<close>
-
-fun pseudo_divmod_main :: "'a :: comm_ring_1  \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly 
-  \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly \<times> 'a poly" where
-  "pseudo_divmod_main lc q r d dr (Suc n) = (let
-     rr = smult lc r;
-     qq = coeff r dr;
-     rrr = rr - monom qq n * d;
-     qqq = smult lc q + monom qq n
-     in pseudo_divmod_main lc qqq rrr d (dr - 1) n)"
-| "pseudo_divmod_main lc q r d dr 0 = (q,r)"
-
-definition pseudo_divmod :: "'a :: comm_ring_1 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" where
-  "pseudo_divmod p q \<equiv> if q = 0 then (0,p) else
-     pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))"
-
-lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \<Longrightarrow> degree rrr \<le> dr \<Longrightarrow> degree rrr \<le> dr - 1"
-  using eq_zero_or_degree_less by fastforce
-  
-lemma pseudo_divmod_main: assumes d: "d \<noteq> 0" "lc = coeff d (degree d)"
-  and *: "degree r \<le> dr" "pseudo_divmod_main lc q r d dr n = (q',r')" 
-    "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> r = 0" 
-  shows "(r' = 0 \<or> degree r' < degree d) \<and> smult (lc^n) (d * q + r) = d * q' + r'"
-  using *
-proof (induct n arbitrary: q r dr)
-  case (Suc n q r dr)
-  let ?rr = "smult lc r"
-  let ?qq = "coeff r dr"
-  define b where [simp]: "b = monom ?qq n"
-  let ?rrr = "?rr - b * d"
-  let ?qqq = "smult lc q + b"
-  note res = Suc(3)
-  from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] 
-  have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" 
-    by (simp del: pseudo_divmod_main.simps)
-  have dr: "dr = n + degree d" using Suc(4) by auto
-  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
-  proof (cases "?qq = 0")
-    case False
-    hence n: "n = degree b" by (simp add: degree_monom_eq)
-    show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum)
-  qed auto
-  also have "\<dots> = lc * coeff b n" unfolding d by simp
-  finally have "coeff (b * d) dr = lc * coeff b n" .
-  moreover have "coeff ?rr dr = lc * coeff r dr" by simp
-  ultimately have c0: "coeff ?rrr dr = 0" by auto
-  have dr: "dr = n + degree d" using Suc(4) by auto
-  have deg_rr: "degree ?rr \<le> dr" using Suc(2) 
-    using degree_smult_le dual_order.trans by blast 
-  have deg_bd: "degree (b * d) \<le> dr"
-    unfolding dr
-    by(rule order.trans[OF degree_mult_le], auto simp: degree_monom_le)
-  have "degree ?rrr \<le> dr"
-    using degree_diff_le[OF deg_rr deg_bd] by auto
-  with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)" by (rule coeff_0_degree_minus_1)
-  have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0"
-  proof (cases dr)
-    case 0
-    with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto
-    with deg_rrr have "degree ?rrr = 0" by simp
-    hence "\<exists> a. ?rrr = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases)
-    from this obtain a where rrr: "?rrr = [:a:]" by auto
-    show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp
-  qed (insert Suc(4), auto)
-  note IH = Suc(1)[OF deg_rrr res this]
-  show ?case
-  proof (intro conjI)
-    show "r' = 0 \<or> degree r' < degree d" using IH by blast
-    show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'"
-      unfolding IH[THEN conjunct2,symmetric]
-      by (simp add: field_simps smult_add_right)
-  qed
-qed auto
-
-lemma pseudo_divmod:
-  assumes g: "g \<noteq> 0" and *: "pseudo_divmod f g = (q,r)" 
-  shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A)
-    and "r = 0 \<or> degree r < degree g" (is ?B)
-proof -
-  from *[unfolded pseudo_divmod_def Let_def]
-  have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g)
-  note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl]
-  have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \<or>
-    degree f = 0 \<and> 1 + length (coeffs f) - length (coeffs g) = 0 \<and> f = 0" using g 
-    by (cases "f = 0"; cases "coeffs g", auto simp: degree_eq_length_coeffs)
-  note main = main[OF this]
-  from main show "r = 0 \<or> degree r < degree g" by auto
-  show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" 
-    by (subst main[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs,
-    insert g, cases "f = 0"; cases "coeffs g", auto)
-qed
-  
-definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)"
-
-lemma snd_pseudo_divmod_main:
-  "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)"
-by (induct n arbitrary: q q' lc r d dr; simp add: Let_def)
-
-definition pseudo_mod 
-    :: "'a :: {comm_ring_1,semiring_1_no_zero_divisors} poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  "pseudo_mod f g = snd (pseudo_divmod f g)"
-  
-lemma pseudo_mod:
-  fixes f g
-  defines "r \<equiv> pseudo_mod f g"
-  assumes g: "g \<noteq> 0"
-  shows "\<exists> a q. a \<noteq> 0 \<and> smult a f = g * q + r" "r = 0 \<or> degree r < degree g"
-proof - 
-  let ?cg = "coeff g (degree g)"
-  let ?cge = "?cg ^ (Suc (degree f) - degree g)"
-  define a where "a = ?cge"
-  obtain q where pdm: "pseudo_divmod f g = (q,r)" using r_def[unfolded pseudo_mod_def]
-    by (cases "pseudo_divmod f g", auto)
-  from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \<or> degree r < degree g" 
-    unfolding a_def by auto
-  show "r = 0 \<or> degree r < degree g" by fact
-  from g have "a \<noteq> 0" unfolding a_def by auto
-  thus "\<exists> a q. a \<noteq> 0 \<and> smult a f = g * q + r" using id by auto
-qed
-
-instantiation poly :: (idom_divide) idom_divide
-begin
-
-fun divide_poly_main :: "'a \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly 
-  \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly" where
-  "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in 
-     if False \<or> a * lc = cr then (* False \<or> is only because of problem in function-package *)
-     divide_poly_main 
-       lc 
-       (q + mon) 
-       (r - mon * d) 
-       d (dr - 1) n else 0)"
-| "divide_poly_main lc q r d dr 0 = q"
-
-definition divide_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  "divide_poly f g = (if g = 0 then 0 else
-     divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" 
-
-lemma divide_poly_main:
-  assumes d: "d \<noteq> 0" "lc = coeff d (degree d)"
-    and *: "degree (d * r) \<le> dr" "divide_poly_main lc q (d * r) d dr n = q'" 
-    "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> d * r = 0" 
-  shows "q' = q + r"
-  using *
-proof (induct n arbitrary: q r dr)
-  case (Suc n q r dr)
-  let ?rr = "d * r"
-  let ?a = "coeff ?rr dr"
-  let ?qq = "?a div lc"
-  define b where [simp]: "b = monom ?qq n"
-  let ?rrr =  "d * (r - b)"
-  let ?qqq = "q + b"
-  note res = Suc(3)
-  have dr: "dr = n + degree d" using Suc(4) by auto
-  have lc: "lc \<noteq> 0" using d by auto
-  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
-  proof (cases "?qq = 0")
-    case False
-    hence n: "n = degree b" by (simp add: degree_monom_eq)
-    show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum)
-  qed simp
-  also have "\<dots> = lc * coeff b n" unfolding d by simp
-  finally have c2: "coeff (b * d) dr = lc * coeff b n" .
-  have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps)
-  have c1: "coeff (d * r) dr = lc * coeff r n"
-  proof (cases "degree r = n")
-    case True
-    thus ?thesis using Suc(2) unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps)
-  next
-    case False
-    have "degree r \<le> n" using dr Suc(2) by auto
-      (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases)
-    with False have r_n: "degree r < n" by auto
-    hence right: "lc * coeff r n = 0" by (simp add: coeff_eq_0)
-    have "coeff (d * r) dr = coeff (d * r) (degree d + n)" unfolding dr by (simp add: ac_simps)
-    also have "\<dots> = 0" using r_n
-      by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 
-        coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq)
-    finally show ?thesis unfolding right .
-  qed
-  have c0: "coeff ?rrr dr = 0" 
-    and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2
-    unfolding b_def coeff_monom coeff_smult c1 using lc by auto
-  from res[unfolded divide_poly_main.simps[of lc q] Let_def] id
-  have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" 
-    by (simp del: divide_poly_main.simps add: field_simps)
-  note IH = Suc(1)[OF _ res] 
-  have dr: "dr = n + degree d" using Suc(4) by auto
-  have deg_rr: "degree ?rr \<le> dr" using Suc(2) by auto
-  have deg_bd: "degree (b * d) \<le> dr"
-    unfolding dr b_def by (rule order.trans[OF degree_mult_le], auto simp: degree_monom_le)
-  have "degree ?rrr \<le> dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd])
-  with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)" by (rule coeff_0_degree_minus_1)
-  have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0"  
-  proof (cases dr)
-    case 0
-    with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto
-    with deg_rrr have "degree ?rrr = 0" by simp
-    from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis
-    show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp
-  qed (insert Suc(4), auto)
-  note IH = IH[OF deg_rrr this]
-  show ?case using IH by simp
-next
-  case (0 q r dr)
-  show ?case 
-  proof (cases "r = 0")
-    case True
-    thus ?thesis using 0 by auto
-  next
-    case False
-    have "degree (d * r) = degree d + degree r" using d False 
-      by (subst degree_mult_eq, auto)
-    thus ?thesis using 0 d by auto
-  qed
-qed 
-
-lemma fst_pseudo_divmod_main_as_divide_poly_main:
-  assumes d: "d \<noteq> 0"
-  defines lc: "lc \<equiv> coeff d (degree d)"
-  shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n"
-proof(induct n arbitrary: q r dr)
-  case 0 then show ?case by simp
-next
-  case (Suc n)
-    note lc0 = leading_coeff_neq_0[OF d, folded lc]
-    then have "pseudo_divmod_main lc q r d dr (Suc n) =
-    pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n)
-      (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n"
-    by (simp add: Let_def ac_simps)
-    also have "fst ... = divide_poly_main lc
-     (smult (lc^n) (smult lc q + monom (coeff r dr) n))
-     (smult (lc^n) (smult lc r - monom (coeff r dr) n * d))
-     d (dr - 1) n"
-      unfolding Suc[unfolded divide_poly_main.simps Let_def]..
-    also have "... = divide_poly_main lc (smult (lc ^ Suc n) q)
-        (smult (lc ^ Suc n) r) d dr (Suc n)"
-      unfolding smult_monom smult_distribs mult_smult_left[symmetric]
-      using lc0 by (simp add: Let_def ac_simps)
-    finally show ?case.
-qed
-
-
-lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0"
-proof (induct n arbitrary: r d dr)
-  case (Suc n r d dr)
-  show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def
-    by (simp add: Suc del: divide_poly_main.simps)
-qed simp
-
-lemma divide_poly:
-  assumes g: "g \<noteq> 0"
-  shows "(f * g) div g = (f :: 'a poly)" 
-proof - 
-  have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs  g)) 
-    = (f * g) div g" unfolding divide_poly_def Let_def by (simp add: ac_simps)
-  note main = divide_poly_main[OF g refl le_refl this]
-  {
-    fix f :: "'a poly"
-    assume "f \<noteq> 0"
-    hence "length (coeffs f) = Suc (degree f)" unfolding degree_eq_length_coeffs by auto
-  } note len = this
-  have "(f * g) div g = 0 + f"
-  proof (rule main, goal_cases)
-    case 1
-    show ?case
-    proof (cases "f = 0")
-      case True
-      with g show ?thesis by (auto simp: degree_eq_length_coeffs)
-    next
-      case False
-      with g have fg: "g * f \<noteq> 0" by auto
-      show ?thesis unfolding len[OF fg] len[OF g] by auto
-    qed
-  qed
-  thus ?thesis by simp
-qed
-
-lemma divide_poly_0: "f div 0 = (0 :: 'a poly)"
-  by (simp add: divide_poly_def Let_def divide_poly_main_0)
-
-instance
-  by standard (auto simp: divide_poly divide_poly_0)
-
-end
-
-instance poly :: (idom_divide) algebraic_semidom ..
-
-lemma div_const_poly_conv_map_poly: 
-  assumes "[:c:] dvd p"
-  shows   "p div [:c:] = map_poly (\<lambda>x. x div c) p"
-proof (cases "c = 0")
-  case False
-  from assms obtain q where p: "p = [:c:] * q" by (erule dvdE)
-  moreover {
-    have "smult c q = [:c:] * q" by simp
-    also have "\<dots> div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto)
-    finally have "smult c q div [:c:] = q" .
-  }
-  ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
-qed (auto intro!: poly_eqI simp: coeff_map_poly)
-
-lemma is_unit_monom_0:
-  fixes a :: "'a::field"
-  assumes "a \<noteq> 0"
-  shows "is_unit (monom a 0)"
-proof
-  from assms show "1 = monom a 0 * monom (inverse a) 0"
-    by (simp add: mult_monom)
-qed
-
-lemma is_unit_triv:
-  fixes a :: "'a::field"
-  assumes "a \<noteq> 0"
-  shows "is_unit [:a:]"
-  using assms by (simp add: is_unit_monom_0 monom_0 [symmetric])
-
-lemma is_unit_iff_degree:
-  assumes "p \<noteq> (0 :: _ :: field poly)"
-  shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q")
-proof
-  assume ?Q
-  then obtain a where "p = [:a:]" by (rule degree_eq_zeroE)
-  with assms show ?P by (simp add: is_unit_triv)
-next
-  assume ?P
-  then obtain q where "q \<noteq> 0" "p * q = 1" ..
-  then have "degree (p * q) = degree 1"
-    by simp
-  with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0"
-    by (simp add: degree_mult_eq)
-  then show ?Q by simp
-qed
-
-lemma is_unit_pCons_iff:
-  "is_unit (pCons (a::_::field) p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
-  by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree)
-
-lemma is_unit_monom_trival:
-  fixes p :: "'a::field poly"
-  assumes "is_unit p"
-  shows "monom (coeff p (degree p)) 0 = p"
-  using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
-
-lemma is_unit_const_poly_iff: 
-  "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \<longleftrightarrow> c dvd 1"
-  by (auto simp: one_poly_def)
-
-lemma is_unit_polyE:
-  fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
-  assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1"
-proof -
-  from assms obtain q where "1 = p * q"
-    by (rule dvdE)
-  then have "p \<noteq> 0" and "q \<noteq> 0"
-    by auto
-  from \<open>1 = p * q\<close> have "degree 1 = degree (p * q)"
-    by simp
-  also from \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> have "\<dots> = degree p + degree q"
-    by (simp add: degree_mult_eq)
-  finally have "degree p = 0" by simp
-  with degree_eq_zeroE obtain c where c: "p = [:c:]" .
-  moreover with \<open>p dvd 1\<close> have "c dvd 1"
-    by (simp add: is_unit_const_poly_iff)
-  ultimately show thesis
-    by (rule that)
-qed
-
-lemma is_unit_polyE':
-  assumes "is_unit (p::_::field poly)"
-  obtains a where "p = monom a 0" and "a \<noteq> 0"
-proof -
-  obtain a q where "p = pCons a q" by (cases p)
-  with assms have "p = [:a:]" and "a \<noteq> 0"
-    by (simp_all add: is_unit_pCons_iff)
-  with that show thesis by (simp add: monom_0)
-qed
-
-lemma is_unit_poly_iff:
-  fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
-  shows "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
-  by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff)
-
-instantiation poly :: ("{normalization_semidom, idom_divide}") normalization_semidom
-begin
-
-definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
-  where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0"
-
-definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
-  where "normalize_poly p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
-
-instance proof
-  fix p :: "'a poly"
-  show "unit_factor p * normalize p = p"
-    by (cases "p = 0")
-       (simp_all add: unit_factor_poly_def normalize_poly_def monom_0 
-          smult_conv_map_poly map_poly_map_poly o_def)
-next
-  fix p :: "'a poly"
-  assume "is_unit p"
-  then obtain c where p: "p = [:c:]" "is_unit c"
-    by (auto simp: is_unit_poly_iff)
-  thus "normalize p = 1"
-    by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def)
-next
-  fix p :: "'a poly" assume "p \<noteq> 0"
-  thus "is_unit (unit_factor p)"
-    by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff)
-qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
-
-end
-
-lemma normalize_poly_eq_div:
-  "normalize p = p div [:unit_factor (lead_coeff p):]"
-proof (cases "p = 0")
-  case False
-  thus ?thesis
-    by (subst div_const_poly_conv_map_poly)
-       (auto simp: normalize_poly_def const_poly_dvd_iff lead_coeff_def )
-qed (auto simp: normalize_poly_def)
-
-lemma unit_factor_pCons:
-  "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)"
-  by (simp add: unit_factor_poly_def)
-
-lemma normalize_monom [simp]:
-  "normalize (monom a n) = monom (normalize a) n"
-  by (simp add: map_poly_monom normalize_poly_def)
-
-lemma unit_factor_monom [simp]:
-  "unit_factor (monom a n) = monom (unit_factor a) 0"
-  by (simp add: unit_factor_poly_def )
-
-lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]"
-  by (simp add: normalize_poly_def map_poly_pCons)
-
-lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)"
-proof -
-  have "smult c p = [:c:] * p" by simp
-  also have "normalize \<dots> = smult (normalize c) (normalize p)"
-    by (subst normalize_mult) (simp add: normalize_const_poly)
-  finally show ?thesis .
-qed
-
-
-subsubsection \<open>Division in Field Polynomials\<close>
-
-text\<open>
- This part connects the above result to the division of field polynomials.
- Mainly imported from Isabelle 2016.
-\<close>
-
-lemma pseudo_divmod_field:
-  assumes g: "(g::'a::field poly) \<noteq> 0" and *: "pseudo_divmod f g = (q,r)"
-  defines "c \<equiv> coeff g (degree g) ^ (Suc (degree f) - degree g)"
-  shows "f = g * smult (1/c) q + smult (1/c) r"
-proof -
-  from leading_coeff_neq_0[OF g] have c0: "c \<noteq> 0" unfolding c_def by auto
-  from pseudo_divmod(1)[OF g *, folded c_def]
-  have "smult c f = g * q + r" by auto
-  also have "smult (1/c) ... = g * smult (1/c) q + smult (1/c) r" by (simp add: smult_add_right)
-  finally show ?thesis using c0 by auto
-qed
-
-lemma divide_poly_main_field:
-  assumes d: "(d::'a::field poly) \<noteq> 0"
-  defines lc: "lc \<equiv> coeff d (degree d)"
-  shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1/lc)^n) q) (smult ((1/lc)^n) r) d dr n)"
-  unfolding lc
-  by(subst fst_pseudo_divmod_main_as_divide_poly_main, auto simp: d power_one_over)
-
-lemma divide_poly_field:
-  fixes f g :: "'a :: field poly"
-  defines "f' \<equiv> smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f"
-  shows "(f::'a::field poly) div g = fst (pseudo_divmod f' g)"
-proof (cases "g = 0")
-  case True show ?thesis 
-    unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0)
-next
-  case False
-    from leading_coeff_neq_0[OF False] have "degree f' = degree f" unfolding f'_def by auto
-    then show ?thesis
-      using length_coeffs_degree[of f'] length_coeffs_degree[of f]
-      unfolding divide_poly_def pseudo_divmod_def Let_def
-                divide_poly_main_field[OF False]
-                length_coeffs_degree[OF False] 
-                f'_def
-      by force
-qed
-
-instantiation poly :: (field) ring_div
-begin
-
-definition modulo_poly where
-  mod_poly_def: "f mod g \<equiv>
-    if g = 0 then f
-    else pseudo_mod (smult ((1/coeff g (degree g)) ^ (Suc (degree f) - degree g)) f) g"
-
-lemma eucl_rel_poly: "eucl_rel_poly (x::'a::field poly) y (x div y, x mod y)"
-  unfolding eucl_rel_poly_iff
-proof (intro conjI)
-  show "x = x div y * y + x mod y"
-  proof(cases "y = 0")
-    case True show ?thesis by(simp add: True divide_poly_def divide_poly_0 mod_poly_def)
-  next
-    case False
-    then have "pseudo_divmod (smult ((1 / coeff y (degree y)) ^ (Suc (degree x) - degree y)) x) y =
-          (x div y, x mod y)"
-      unfolding divide_poly_field mod_poly_def pseudo_mod_def by simp
-    from pseudo_divmod[OF False this]
-    show ?thesis using False
-      by (simp add: power_mult_distrib[symmetric] mult.commute)
-  qed
-  show "if y = 0 then x div y = 0 else x mod y = 0 \<or> degree (x mod y) < degree y"
-  proof (cases "y = 0")
-    case True then show ?thesis by auto
-  next
-    case False
-      with pseudo_mod[OF this] show ?thesis unfolding mod_poly_def by simp
-  qed
-qed
-
-lemma div_poly_eq:
-  "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x div y = q"
-  by(rule eucl_rel_poly_unique_div[OF eucl_rel_poly])
-
-lemma mod_poly_eq:
-  "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x mod y = r"
-  by (rule eucl_rel_poly_unique_mod[OF eucl_rel_poly])
-
-instance
-proof
-  fix x y :: "'a poly"
-  from eucl_rel_poly[of x y,unfolded eucl_rel_poly_iff]
-  show "x div y * y + x mod y = x" by auto
-next
-  fix x y z :: "'a poly"
-  assume "y \<noteq> 0"
-  hence "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
-    using eucl_rel_poly [of x y]
-    by (simp add: eucl_rel_poly_iff distrib_right)
-  thus "(x + z * y) div y = z + x div y"
-    by (rule div_poly_eq)
-next
-  fix x y z :: "'a poly"
-  assume "x \<noteq> 0"
-  show "(x * y) div (x * z) = y div z"
-  proof (cases "y \<noteq> 0 \<and> z \<noteq> 0")
-    have "\<And>x::'a poly. eucl_rel_poly x 0 (0, x)"
-      by (rule eucl_rel_poly_by_0)
-    then have [simp]: "\<And>x::'a poly. x div 0 = 0"
-      by (rule div_poly_eq)
-    have "\<And>x::'a poly. eucl_rel_poly 0 x (0, 0)"
-      by (rule eucl_rel_poly_0)
-    then have [simp]: "\<And>x::'a poly. 0 div x = 0"
-      by (rule div_poly_eq)
-    case False then show ?thesis by auto
-  next
-    case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto
-    with \<open>x \<noteq> 0\<close>
-    have "\<And>q r. eucl_rel_poly y z (q, r) \<Longrightarrow> eucl_rel_poly (x * y) (x * z) (q, x * r)"
-      by (auto simp add: eucl_rel_poly_iff algebra_simps)
-        (rule classical, simp add: degree_mult_eq)
-    moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" .
-    ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" .
-    then show ?thesis by (simp add: div_poly_eq)
-  qed
-qed
-
-end
-
-lemma degree_mod_less:
-  "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
-  using eucl_rel_poly [of x y]
-  unfolding eucl_rel_poly_iff by simp
-
-lemma div_poly_less: "degree (x::'a::field poly) < degree y \<Longrightarrow> x div y = 0"
-proof -
-  assume "degree x < degree y"
-  hence "eucl_rel_poly x y (0, x)"
-    by (simp add: eucl_rel_poly_iff)
-  thus "x div y = 0" by (rule div_poly_eq)
-qed
-
-lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x"
-proof -
-  assume "degree x < degree y"
-  hence "eucl_rel_poly x y (0, x)"
-    by (simp add: eucl_rel_poly_iff)
-  thus "x mod y = x" by (rule mod_poly_eq)
-qed
-
-lemma eucl_rel_poly_smult_left:
-  "eucl_rel_poly x y (q, r)
-    \<Longrightarrow> eucl_rel_poly (smult a x) y (smult a q, smult a r)"
-  unfolding eucl_rel_poly_iff by (simp add: smult_add_right)
-
-lemma div_smult_left: "(smult (a::'a::field) x) div y = smult a (x div y)"
-  by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
-
-lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)"
-  by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
-
-lemma poly_div_minus_left [simp]:
-  fixes x y :: "'a::field poly"
-  shows "(- x) div y = - (x div y)"
-  using div_smult_left [of "- 1::'a"] by simp
-
-lemma poly_mod_minus_left [simp]:
-  fixes x y :: "'a::field poly"
-  shows "(- x) mod y = - (x mod y)"
-  using mod_smult_left [of "- 1::'a"] by simp
-
-lemma eucl_rel_poly_add_left:
-  assumes "eucl_rel_poly x y (q, r)"
-  assumes "eucl_rel_poly x' y (q', r')"
-  shows "eucl_rel_poly (x + x') y (q + q', r + r')"
-  using assms unfolding eucl_rel_poly_iff
-  by (auto simp add: algebra_simps degree_add_less)
-
-lemma poly_div_add_left:
-  fixes x y z :: "'a::field poly"
-  shows "(x + y) div z = x div z + y div z"
-  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
-  by (rule div_poly_eq)
-
-lemma poly_mod_add_left:
-  fixes x y z :: "'a::field poly"
-  shows "(x + y) mod z = x mod z + y mod z"
-  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
-  by (rule mod_poly_eq)
-
-lemma poly_div_diff_left:
-  fixes x y z :: "'a::field poly"
-  shows "(x - y) div z = x div z - y div z"
-  by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left)
-
-lemma poly_mod_diff_left:
-  fixes x y z :: "'a::field poly"
-  shows "(x - y) mod z = x mod z - y mod z"
-  by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left)
-
-lemma eucl_rel_poly_smult_right:
-  "a \<noteq> 0 \<Longrightarrow> eucl_rel_poly x y (q, r)
-    \<Longrightarrow> eucl_rel_poly x (smult a y) (smult (inverse a) q, r)"
-  unfolding eucl_rel_poly_iff by simp
-
-lemma div_smult_right:
-  "(a::'a::field) \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
-  by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
-
-lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y"
-  by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
-
-lemma poly_div_minus_right [simp]:
-  fixes x y :: "'a::field poly"
-  shows "x div (- y) = - (x div y)"
-  using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq)
-
-lemma poly_mod_minus_right [simp]:
-  fixes x y :: "'a::field poly"
-  shows "x mod (- y) = x mod y"
-  using mod_smult_right [of "- 1::'a"] by simp
-
-lemma eucl_rel_poly_mult:
-  "eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly q z (q', r')
-    \<Longrightarrow> eucl_rel_poly x (y * z) (q', y * r' + r)"
-apply (cases "z = 0", simp add: eucl_rel_poly_iff)
-apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff)
-apply (cases "r = 0")
-apply (cases "r' = 0")
-apply (simp add: eucl_rel_poly_iff)
-apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq)
-apply (cases "r' = 0")
-apply (simp add: eucl_rel_poly_iff degree_mult_eq)
-apply (simp add: eucl_rel_poly_iff field_simps)
-apply (simp add: degree_mult_eq degree_add_less)
-done
-
-lemma poly_div_mult_right:
-  fixes x y z :: "'a::field poly"
-  shows "x div (y * z) = (x div y) div z"
-  by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
-
-lemma poly_mod_mult_right:
-  fixes x y z :: "'a::field poly"
-  shows "x mod (y * z) = y * (x div y mod z) + x mod y"
-  by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
-
-lemma mod_pCons:
-  fixes a and x
-  assumes y: "y \<noteq> 0"
-  defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)"
-  shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)"
-unfolding b
-apply (rule mod_poly_eq)
-apply (rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl])
-done
-
-definition pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly"
-where
-  "pdivmod p q = (p div q, p mod q)"
-
-lemma pdivmod_0:
-  "pdivmod 0 q = (0, 0)"
-  by (simp add: pdivmod_def)
-
-lemma pdivmod_pCons:
-  "pdivmod (pCons a p) q =
-    (if q = 0 then (0, pCons a p) else
-      (let (s, r) = pdivmod p q;
-           b = coeff (pCons a r) (degree q) / coeff q (degree q)
-        in (pCons b s, pCons a r - smult b q)))"
-  apply (simp add: pdivmod_def Let_def, safe)
-  apply (rule div_poly_eq)
-  apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
-  apply (rule mod_poly_eq)
-  apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
-  done
-
-lemma pdivmod_fold_coeffs:
-  "pdivmod p q = (if q = 0 then (0, p)
-    else fold_coeffs (\<lambda>a (s, r).
-      let b = coeff (pCons a r) (degree q) / coeff q (degree q)
-      in (pCons b s, pCons a r - smult b q)
-   ) p (0, 0))"
-  apply (cases "q = 0")
-  apply (simp add: pdivmod_def)
-  apply (rule sym)
-  apply (induct p)
-  apply (simp_all add: pdivmod_0 pdivmod_pCons)
-  apply (case_tac "a = 0 \<and> p = 0")
-  apply (auto simp add: pdivmod_def)
-  done
-
-subsection \<open>List-based versions for fast implementation\<close>
-(* Subsection by:
-      Sebastiaan Joosten
-      René Thiemann
-      Akihisa Yamada
-    *)
-fun minus_poly_rev_list :: "'a :: group_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)"
-| "minus_poly_rev_list xs [] = xs"
-| "minus_poly_rev_list [] (y # ys) = []"
-
-fun pseudo_divmod_main_list :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list 
-  \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" where
-  "pseudo_divmod_main_list lc q r d (Suc n) = (let
-     rr = map (op * lc) r;
-     a = hd r;
-     qqq = cCons a (map (op * lc) q);
-     rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d))
-     in pseudo_divmod_main_list lc qqq rrr d n)"
-| "pseudo_divmod_main_list lc q r d 0 = (q,r)"
-
-fun pseudo_mod_main_list :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list 
-  \<Rightarrow> nat \<Rightarrow> 'a list" where
-  "pseudo_mod_main_list lc r d (Suc n) = (let
-     rr = map (op * lc) r;
-     a = hd r;
-     rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d))
-     in pseudo_mod_main_list lc rrr d n)"
-| "pseudo_mod_main_list lc r d 0 = r"
-
-
-fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list 
-  \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" where
-  "divmod_poly_one_main_list q r d (Suc n) = (let
-     a = hd r;
-     qqq = cCons a q;
-     rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d))
-     in divmod_poly_one_main_list qqq rr d n)"
-| "divmod_poly_one_main_list q r d 0 = (q,r)"
-
-fun mod_poly_one_main_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list 
-  \<Rightarrow> nat \<Rightarrow> 'a list" where
-  "mod_poly_one_main_list r d (Suc n) = (let
-     a = hd r;
-     rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d))
-     in mod_poly_one_main_list rr d n)"
-| "mod_poly_one_main_list r d 0 = r"
-
-definition pseudo_divmod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" where
-  "pseudo_divmod_list p q =
-  (if q = [] then ([],p) else
- (let rq = rev q;
-     (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in 
-   (qu,rev re)))"
-
-definition pseudo_mod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "pseudo_mod_list p q =
-  (if q = [] then p else
- (let rq = rev q;
-     re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in 
-   (rev re)))"
-
-lemma minus_zero_does_nothing:
-"(minus_poly_rev_list x (map (op * 0) y)) = (x :: 'a :: ring list)"
-  by(induct x y rule: minus_poly_rev_list.induct, auto)
-
-lemma length_minus_poly_rev_list[simp]:
- "length (minus_poly_rev_list xs ys) = length xs"
-  by(induct xs ys rule: minus_poly_rev_list.induct, auto)
-
-lemma if_0_minus_poly_rev_list:
-  "(if a = 0 then x else minus_poly_rev_list x (map (op * a) y))
-      = minus_poly_rev_list x (map (op * (a :: 'a :: ring)) y)"
-  by(cases "a=0",simp_all add:minus_zero_does_nothing)
-
-lemma Poly_append:
-  "Poly ((a::'a::comm_semiring_1 list) @ b) = Poly a + monom 1 (length a) * Poly b"
-  by (induct a,auto simp: monom_0 monom_Suc)
-
-lemma minus_poly_rev_list: "length p \<ge> length (q :: 'a :: comm_ring_1 list) \<Longrightarrow>
-  Poly (rev (minus_poly_rev_list (rev p) (rev q)))
-  = Poly p - monom 1 (length p - length q) * Poly q"
-proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct)
-  case (1 x xs y ys) 
-  have "length (rev q) \<le> length (rev p)" using 1 by simp
-  from this[folded 1(2,3)] have ys_xs:"length ys \<le> length xs" by simp
-  hence a:"Poly (rev (minus_poly_rev_list xs ys)) =
-           Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)"
-    by(subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev],auto)
-  have "Poly p - monom 1 (length p - length q) * Poly q
-      = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))"
-    by simp
-  also have "\<dots> = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))"
-    unfolding 1(2,3) by simp
-  also have "\<dots> = Poly (rev xs) + monom x (length xs) -
-  (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" using ys_xs
-    by (simp add:Poly_append distrib_left mult_monom smult_monom)
-  also have "\<dots> = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)"
-    unfolding a diff_monom[symmetric] by(simp)
-  finally show ?case
-    unfolding 1(2,3)[symmetric] by (simp add: smult_monom Poly_append)
-qed auto
-
-lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f"
-  using smult_monom [of a _ n] by (metis mult_smult_left)
-
-lemma head_minus_poly_rev_list:
-  "length d \<le> length r \<Longrightarrow> d\<noteq>[] \<Longrightarrow>
-  hd (minus_poly_rev_list (map (op * (last d :: 'a :: comm_ring)) r) (map (op * (hd r)) (rev d))) = 0"
-proof(induct r)
-  case (Cons a rs)
-  thus ?case by(cases "rev d", simp_all add:ac_simps)
-qed simp
-
-lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)"
-proof (induct p)
-  case(Cons x xs) thus ?case by (cases "Poly xs = 0",auto)
-qed simp
-
-lemma last_coeff_is_hd: "xs \<noteq> [] \<Longrightarrow> coeff (Poly xs) (length xs - 1) = hd (rev xs)"
-  by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append)
-
-lemma pseudo_divmod_main_list_invar :
-  assumes leading_nonzero:"last d \<noteq> 0"
-  and lc:"last d = lc"
-  and dNonempty:"d \<noteq> []"
-  and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q',rev r')"
-  and "n = (1 + length r - length d)"
-  shows 
-  "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = 
-  (Poly q', Poly r')"
-using assms(4-)
-proof(induct "n" arbitrary: r q)
-case (Suc n r q)
-  have ifCond: "\<not> Suc (length r) \<le> length d" using Suc.prems by simp
-  have rNonempty:"r \<noteq> []"
-    using ifCond dNonempty using Suc_leI length_greater_0_conv list.size(3) by fastforce
-  let ?a = "(hd (rev r))"
-  let ?rr = "map (op * lc) (rev r)"
-  let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map (op * ?a) (rev d))))"
-  let ?qq = "cCons ?a (map (op * lc) q)"
-  have n: "n = (1 + length r - length d - 1)"
-    using ifCond Suc(3) by simp
-  have rr_val:"(length ?rrr) = (length r - 1)" using ifCond by auto
-  hence rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)"
-    using rNonempty ifCond unfolding One_nat_def by auto
-  from ifCond have id: "Suc (length r) - length d = Suc (length r - length d)" by auto
-  have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')"
-    using Suc.prems ifCond by (simp add:Let_def if_0_minus_poly_rev_list id)
-  hence v:"pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')"
-    using n by auto
-  have sucrr:"Suc (length r) - length d = Suc (length r - length d)"
-    using Suc_diff_le ifCond not_less_eq_eq by blast
-  have n_ok : "n = 1 + (length ?rrr) - length d" using Suc(3) rNonempty by simp
-  have cong: "\<And> x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow>
-    pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp
-  have hd_rev:"coeff (Poly r) (length r - Suc 0) = hd (rev r)"
-    using last_coeff_is_hd[OF rNonempty] by simp
-  show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def
-  proof (rule cong[OF _ _ refl], goal_cases)
-    case 1 
-    show ?case unfolding monom_Suc hd_rev[symmetric]
-      by (simp add: smult_monom Poly_map)
-  next
-    case 2 
-    show ?case 
-    proof (subst Poly_on_rev_starting_with_0, goal_cases)
-      show "hd (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))) = 0"
-        by (fold lc, subst head_minus_poly_rev_list, insert ifCond dNonempty,auto)
-      from ifCond have "length d \<le> length r" by simp
-      then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d =
-        Poly (rev (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))))"
-        by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric]
-          minus_poly_rev_list)
-    qed
-  qed simp
-qed simp
-
-lemma pseudo_divmod_impl[code]: "pseudo_divmod (f::'a::comm_ring_1 poly) g =
-  map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))"
-proof (cases "g=0")
-case False
-  hence coeffs_g_nonempty:"(coeffs g) \<noteq> []" by simp
-  hence lastcoeffs:"last (coeffs g) = coeff g (degree g)"
-    by (simp add: hd_rev last_coeffs_eq_coeff_degree not_0_coeffs_not_Nil)
-  obtain q r where qr: "pseudo_divmod_main_list
-            (last (coeffs g)) (rev [])
-            (rev (coeffs f)) (rev (coeffs g))
-            (1 + length (coeffs f) -
-             length (coeffs g)) = (q,rev (rev r))"  by force
-  then have qr': "pseudo_divmod_main_list
-            (hd (rev (coeffs g))) []
-            (rev (coeffs f)) (rev (coeffs g))
-            (1 + length (coeffs f) -
-             length (coeffs g)) = (q,r)" using hd_rev[OF coeffs_g_nonempty] by(auto)
-  from False have cg: "(coeffs g = []) = False" by auto
-  have last_non0:"last (coeffs g) \<noteq> 0" using False by (simp add:last_coeffs_not_0)
-  show ?thesis
-    unfolding pseudo_divmod_def pseudo_divmod_list_def Let_def qr' map_prod_def split cg if_False
-    pseudo_divmod_main_list_invar[OF last_non0 _ _ qr,unfolded lastcoeffs,simplified,symmetric,OF False]
-    poly_of_list_def
-    using False by (auto simp: degree_eq_length_coeffs)
-next
-  case True
-  show ?thesis unfolding True unfolding pseudo_divmod_def pseudo_divmod_list_def
-  by auto
-qed
-
-lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q 
-  xs ys n) = pseudo_mod_main_list l xs ys n"
-  by (induct n arbitrary: l q xs ys, auto simp: Let_def)
-
-lemma pseudo_mod_impl[code]: "pseudo_mod f g =
-  poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))"
-proof -
-  have snd_case: "\<And> f g p. snd ((\<lambda> (x,y). (f x, g y)) p) = g (snd p)" 
-    by auto
-  show ?thesis
-  unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def
-    pseudo_mod_list_def Let_def
-  by (simp add: snd_case pseudo_mod_main_list)
-qed
-
-
-(* *************** *)
-subsubsection \<open>Improved Code-Equations for Polynomial (Pseudo) Division\<close>
-
-lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \<longleftrightarrow> pdivmod p q = (r, s)"
-  by (metis pdivmod_def eucl_rel_poly eucl_rel_poly_unique)
-
-lemma pdivmod_via_pseudo_divmod: "pdivmod f g = (if g = 0 then (0,f) 
-     else let 
-       ilc = inverse (coeff g (degree g));       
-       h = smult ilc g;
-       (q,r) = pseudo_divmod f h
-     in (smult ilc q, r))" (is "?l = ?r")
-proof (cases "g = 0")
-  case False
-  define lc where "lc = inverse (coeff g (degree g))"
-  define h where "h = smult lc g"
-  from False have h1: "coeff h (degree h) = 1" and lc: "lc \<noteq> 0" unfolding h_def lc_def by auto
-  hence h0: "h \<noteq> 0" by auto
-  obtain q r where p: "pseudo_divmod f h = (q,r)" by force
-  from False have id: "?r = (smult lc q, r)" 
-    unfolding Let_def h_def[symmetric] lc_def[symmetric] p by auto
-  from pseudo_divmod[OF h0 p, unfolded h1] 
-  have f: "f = h * q + r" and r: "r = 0 \<or> degree r < degree h" by auto
-  have "eucl_rel_poly f h (q, r)" unfolding eucl_rel_poly_iff using f r h0 by auto
-  hence "pdivmod f h = (q,r)" by (simp add: pdivmod_pdivmodrel)
-  hence "pdivmod f g = (smult lc q, r)" 
-    unfolding pdivmod_def h_def div_smult_right[OF lc] mod_smult_right[OF lc]
-    using lc by auto
-  with id show ?thesis by auto
-qed (auto simp: pdivmod_def)
-
-lemma pdivmod_via_pseudo_divmod_list: "pdivmod f g = (let 
-  cg = coeffs g
-  in if cg = [] then (0,f)
-     else let 
-       cf = coeffs f;
-       ilc = inverse (last cg);       
-       ch = map (op * ilc) cg;
-       (q,r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg)
-     in (poly_of_list (map (op * ilc) q), poly_of_list (rev r)))"
-proof -
-  note d = pdivmod_via_pseudo_divmod
-      pseudo_divmod_impl pseudo_divmod_list_def
-  show ?thesis
-  proof (cases "g = 0")
-    case True thus ?thesis unfolding d by auto
-  next
-    case False
-    define ilc where "ilc = inverse (coeff g (degree g))"
-    from False have ilc: "ilc \<noteq> 0" unfolding ilc_def by auto
-    with False have id: "(g = 0) = False" "(coeffs g = []) = False" 
-      "last (coeffs g) = coeff g (degree g)" 
-      "(coeffs (smult ilc g) = []) = False"
-      by (auto simp: last_coeffs_eq_coeff_degree) 
-    have id2: "hd (rev (coeffs (smult ilc g))) = 1"      
-      by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def)
-    have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" 
-      "rev (coeffs (smult ilc g)) = rev (map (op * ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto
-    obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map (op * ilc) (coeffs g)))
-           (1 + length (coeffs f) - length (coeffs g)) = (q,r)" by force
-    show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 
-      unfolding id3 pair map_prod_def split by (auto simp: Poly_map)
-  qed
-qed
-
-lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list"
-proof (intro ext, goal_cases)
-  case (1 q r d n)
-  {
-    fix xs :: "'a list"
-    have "map (op * 1) xs = xs" by (induct xs, auto)
-  } note [simp] = this
-  show ?case by (induct n arbitrary: q r d, auto simp: Let_def)
-qed
-
-fun divide_poly_main_list :: "'a::idom_divide \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list 
-  \<Rightarrow> nat \<Rightarrow> 'a list" where
-  "divide_poly_main_list lc q r d (Suc n) = (let
-     cr = hd r
-     in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let 
-     a = cr div lc;
-     qq = cCons a q;
-     rr = minus_poly_rev_list r (map (op * a) d)
-     in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])"
-| "divide_poly_main_list lc q r d 0 = q"
-
-
-lemma divide_poly_main_list_simp[simp]: "divide_poly_main_list lc q r d (Suc n) = (let
-     cr = hd r;
-     a = cr div lc;
-     qq = cCons a q;
-     rr = minus_poly_rev_list r (map (op * a) d)
-     in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])"
-  by (simp add: Let_def minus_zero_does_nothing)
-
-declare divide_poly_main_list.simps(1)[simp del]
-
-definition divide_poly_list :: "'a::idom_divide poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
-  "divide_poly_list f g =
-    (let cg = coeffs g
-     in if cg = [] then g
-        else let cf = coeffs f; cgr = rev cg
-          in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))"
-
-lemmas pdivmod_via_divmod_list[code] = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1]
-
-lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n"
-  by  (induct n arbitrary: q r d, auto simp: Let_def)
-
-lemma mod_poly_code[code]: "f mod g =
-    (let cg = coeffs g
-     in if cg = [] then f
-        else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg;
-                 r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg)
-             in poly_of_list (rev r))" (is "?l = ?r")
-proof -
-  have "?l = snd (pdivmod f g)" unfolding pdivmod_def by simp
-  also have "\<dots> = ?r" unfolding pdivmod_via_divmod_list Let_def
-     mod_poly_one_main_list[symmetric, of _ _ _ Nil] by (auto split: prod.splits)
-  finally show ?thesis .
-qed
-
-definition div_field_poly_impl :: "'a :: field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where 
-  "div_field_poly_impl f g = (
-    let cg = coeffs g
-      in if cg = [] then 0
-        else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg;
-                 q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg))
-             in poly_of_list ((map (op * ilc) q)))"
-
-text \<open>We do not declare the following lemma as code equation, since then polynomial division 
-  on non-fields will no longer be executable. However, a code-unfold is possible, since 
-  \<open>div_field_poly_impl\<close> is a bit more efficient than the generic polynomial division.\<close>
-lemma div_field_poly_impl[code_unfold]: "op div = div_field_poly_impl"
-proof (intro ext)
-  fix f g :: "'a poly"
-  have "f div g = fst (pdivmod f g)" unfolding pdivmod_def by simp
-  also have "\<dots> = div_field_poly_impl f g" unfolding 
-    div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits)
-  finally show "f div g =  div_field_poly_impl f g" .
-qed
-
-
-lemma divide_poly_main_list:
-  assumes lc0: "lc \<noteq> 0"
-  and lc:"last d = lc"
-  and d:"d \<noteq> []"
-  and "n = (1 + length r - length d)"
-  shows 
-  "Poly (divide_poly_main_list lc q (rev r) (rev d) n) =
-  divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n"
-using assms(4-)
-proof(induct "n" arbitrary: r q)
-case (Suc n r q)
-  have ifCond: "\<not> Suc (length r) \<le> length d" using Suc.prems by simp
-  have r: "r \<noteq> []"
-    using ifCond d using Suc_leI length_greater_0_conv list.size(3) by fastforce
-  then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases, auto)
-  from d lc obtain dd where d: "d = dd @ [lc]" 
-    by (cases d rule: rev_cases, auto)
-  from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r)
-  from ifCond have len: "length dd \<le> length rr" by (simp add: r d)
-  show ?case
-  proof (cases "lcr div lc * lc = lcr")
-    case False
-    thus ?thesis unfolding Suc(2)[symmetric] using r d
-      by (auto simp add: Let_def nth_default_append)
-  next
-    case True
-    hence id:
-    "?thesis = (Poly (divide_poly_main_list lc (cCons (lcr div lc) q)
-         (rev (rev (minus_poly_rev_list (rev rr) (rev (map (op * (lcr div lc)) dd))))) (rev d) n) = 
-      divide_poly_main lc
-           (monom 1 (Suc n) * Poly q + monom (lcr div lc) n)
-           (Poly r - monom (lcr div lc) n * Poly d)
-           (Poly d) (length rr - 1) n)"
-           using r d 
-      by (cases r rule: rev_cases; cases "d" rule: rev_cases; 
-        auto simp add: Let_def rev_map nth_default_append)      
-    have cong: "\<And> x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow>
-      divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp
-    show ?thesis unfolding id 
-    proof (subst Suc(1), simp add: n,
-      subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases)
-      case 2 
-      have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)"
-        by (simp add: mult_monom len True)
-      thus ?case unfolding r d Poly_append n ring_distribs
-        by (auto simp: Poly_map smult_monom smult_monom_mult)
-    qed (auto simp: len monom_Suc smult_monom)
-  qed
-qed simp
-
-
-lemma divide_poly_list[code]: "f div g = divide_poly_list f g" 
-proof -
-  note d = divide_poly_def divide_poly_list_def
-  show ?thesis
-  proof (cases "g = 0")
-    case True
-    show ?thesis unfolding d True by auto
-  next
-    case False
-    then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases, auto)    
-    with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto
-    from cg False have lcg: "coeff g (degree g) = lcg" 
-      using last_coeffs_eq_coeff_degree last_snoc by force
-    with False have lcg0: "lcg \<noteq> 0" by auto
-    from cg have ltp: "Poly (cg @ [lcg]) = g"
-     using Poly_coeffs [of g] by auto
-    show ?thesis unfolding d cg Let_def id if_False poly_of_list_def
-      by (subst divide_poly_main_list, insert False cg lcg0, auto simp: lcg ltp,
-      simp add: degree_eq_length_coeffs)
-  qed
-qed
-
-subsection \<open>Order of polynomial roots\<close>
+    
+subsubsection \<open>Order of polynomial roots\<close>
 
 definition order :: "'a::idom \<Rightarrow> 'a poly \<Rightarrow> nat"
 where
@@ -2986,6 +1787,124 @@
 lemma order_0I: "poly p a \<noteq> 0 \<Longrightarrow> order a p = 0"
   by (subst (asm) order_root) auto
 
+lemma order_unique_lemma:
+  fixes p :: "'a::idom poly"
+  assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-a, 1:] ^ Suc n dvd p"
+  shows "n = order a p"
+unfolding Polynomial.order_def
+apply (rule Least_equality [symmetric])
+apply (fact assms)
+apply (rule classical)
+apply (erule notE)
+unfolding not_less_eq_eq
+using assms(1) apply (rule power_le_dvd)
+apply assumption
+  done
+    
+lemma order_mult: "p * q \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"
+proof -
+  define i where "i = order a p"
+  define j where "j = order a q"
+  define t where "t = [:-a, 1:]"
+  have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"
+    unfolding t_def by (simp add: dvd_iff_poly_eq_0)
+  assume "p * q \<noteq> 0"
+  then show "order a (p * q) = i + j"
+    apply clarsimp
+    apply (drule order [where a=a and p=p, folded i_def t_def])
+    apply (drule order [where a=a and p=q, folded j_def t_def])
+    apply clarify
+    apply (erule dvdE)+
+    apply (rule order_unique_lemma [symmetric], fold t_def)
+    apply (simp_all add: power_add t_dvd_iff)
+    done
+qed
+
+lemma order_smult:
+  assumes "c \<noteq> 0" 
+  shows "order x (smult c p) = order x p"
+proof (cases "p = 0")
+  case False
+  have "smult c p = [:c:] * p" by simp
+  also from assms False have "order x \<dots> = order x [:c:] + order x p" 
+    by (subst order_mult) simp_all
+  also from assms have "order x [:c:] = 0" by (intro order_0I) auto
+  finally show ?thesis by simp
+qed simp
+
+(* Next two lemmas contributed by Wenda Li *)
+lemma order_1_eq_0 [simp]:"order x 1 = 0" 
+  by (metis order_root poly_1 zero_neq_one)
+
+lemma order_power_n_n: "order a ([:-a,1:]^n)=n" 
+proof (induct n) (*might be proved more concisely using nat_less_induct*)
+  case 0
+  thus ?case by (metis order_root poly_1 power_0 zero_neq_one)
+next 
+  case (Suc n)
+  have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" 
+    by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral 
+      one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right)
+  moreover have "order a [:-a,1:]=1" unfolding order_def
+    proof (rule Least_equality,rule ccontr)
+      assume  "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
+      hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp
+      hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )" 
+        by (rule dvd_imp_degree_le,auto) 
+      thus False by auto
+    next
+      fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
+      show "1 \<le> y" 
+        proof (rule ccontr)
+          assume "\<not> 1 \<le> y"
+          hence "y=0" by auto
+          hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto
+          thus False using asm by auto
+        qed
+    qed
+  ultimately show ?case using Suc by auto
+qed
+
+lemma order_0_monom [simp]:
+  assumes "c \<noteq> 0"
+  shows   "order 0 (monom c n) = n"
+  using assms order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult)
+
+lemma dvd_imp_order_le:
+  "q \<noteq> 0 \<Longrightarrow> p dvd q \<Longrightarrow> Polynomial.order a p \<le> Polynomial.order a q"
+  by (auto simp: order_mult elim: dvdE)
+
+text\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close>
+
+lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"
+apply (cases "p = 0", auto)
+apply (drule order_2 [where a=a and p=p])
+apply (metis not_less_eq_eq power_le_dvd)
+apply (erule power_le_dvd [OF order_1])
+done
+
+lemma order_decomp:
+  assumes "p \<noteq> 0"
+  shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q"
+proof -
+  from assms have A: "[:- a, 1:] ^ order a p dvd p"
+    and B: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order)
+  from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" ..
+  with B have "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q"
+    by simp
+  then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q"
+    by simp
+  then have D: "\<not> [:- a, 1:] dvd q"
+    using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q]
+    by auto
+  from C D show ?thesis by blast
+qed
+
+lemma monom_1_dvd_iff:
+  assumes "p \<noteq> 0"
+  shows   "monom 1 n dvd p \<longleftrightarrow> n \<le> order 0 p"
+  using assms order_divides[of 0 n p] by (simp add: monom_altdef)
+
 
 subsection \<open>Additional induction rules on polynomials\<close>
 
@@ -3055,7 +1974,7 @@
   finally show ?thesis .
 qed
 
-
+  
 subsection \<open>Composition of polynomials\<close>
 
 (* Several lemmas contributed by René Thiemann and Akihisa Yamada *)
@@ -3218,7 +2137,7 @@
   shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)"
 proof (induct p)
   case 0
-  thus ?case unfolding lead_coeff_def by auto
+  thus ?case by auto
 next
   case (pCons a p)
   have "degree ( q * pcompose p q) = 0 \<Longrightarrow> ?case"
@@ -3230,9 +2149,13 @@
     qed
   moreover have "degree ( q * pcompose p q) > 0 \<Longrightarrow> ?case" 
     proof -
-      assume "degree ( q * pcompose p q) > 0"
-      hence "lead_coeff (pcompose (pCons a p) q) =lead_coeff ( q * pcompose p q)"
-        by (auto simp add:pcompose_pCons lead_coeff_add_le)
+      assume "degree (q * pcompose p q) > 0"
+      then have "degree [:a:] < degree (q * pcompose p q)"
+        by simp
+      then have "lead_coeff ([:a:] + q * p \<circ>\<^sub>p q) = lead_coeff (q * p \<circ>\<^sub>p q)"
+        by (rule lead_coeff_add_le)
+      then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)"
+        by (simp add: pcompose_pCons)
       also have "... = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)"
         using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp
       also have "... = lead_coeff p * lead_coeff q ^ (degree p + 1)"
@@ -3254,7 +2177,6 @@
 lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)"
   by (auto simp add: nth_default_def add_ac)
   
-  
 lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)"
 proof -
   from MOST_coeff_eq_0[of p] obtain m where "\<forall>k>m. coeff p k = 0" by (auto simp: MOST_nat)
@@ -3359,7 +2281,7 @@
   by (induction rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom)
 
 lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p"
-  by (simp add: coeff_reflect_poly lead_coeff_def)
+  by (simp add: coeff_reflect_poly)
 
 lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p"
   by (simp add: poly_0_coeff_0)
@@ -3442,7 +2364,7 @@
   reflect_poly_power reflect_poly_prod reflect_poly_prod_list
 
 
-subsection \<open>Derivatives of univariate polynomials\<close>
+subsection \<open>Derivatives\<close>
 
 function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \<Rightarrow> 'a poly"
 where
@@ -3735,6 +2657,136 @@
   qed
 qed
 
+lemma lemma_order_pderiv1:
+  "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q +
+    smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)"
+apply (simp only: pderiv_mult pderiv_power_Suc)
+apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons)
+done
+
+lemma lemma_order_pderiv:
+  fixes p :: "'a :: field_char_0 poly"
+  assumes n: "0 < n" 
+      and pd: "pderiv p \<noteq> 0" 
+      and pe: "p = [:- a, 1:] ^ n * q" 
+      and nd: "~ [:- a, 1:] dvd q"
+    shows "n = Suc (order a (pderiv p))"
+using n 
+proof -
+  have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0"
+    using assms by auto
+  obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
+    using assms by (cases n) auto
+  have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l"
+    by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff)
+  have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" 
+  proof (rule order_unique_lemma)
+    show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
+      apply (subst lemma_order_pderiv1)
+      apply (rule dvd_add)
+      apply (metis dvdI dvd_mult2 power_Suc2)
+      apply (metis dvd_smult dvd_triv_right)
+      done
+  next
+    show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
+     apply (subst lemma_order_pderiv1)
+     by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one)
+  qed
+  then show ?thesis
+    by (metis \<open>n = Suc n'\<close> pe)
+qed
+
+lemma order_pderiv:
+  "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> \<Longrightarrow>
+     (order a p = Suc (order a (pderiv p)))"
+apply (case_tac "p = 0", simp)
+apply (drule_tac a = a and p = p in order_decomp)
+using neq0_conv
+apply (blast intro: lemma_order_pderiv)
+done
+
+lemma poly_squarefree_decomp_order:
+  assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0"
+  and p: "p = q * d"
+  and p': "pderiv p = e * d"
+  and d: "d = r * p + s * pderiv p"
+  shows "order a q = (if order a p = 0 then 0 else 1)"
+proof (rule classical)
+  assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)"
+  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
+  with p have "order a p = order a q + order a d"
+    by (simp add: order_mult)
+  with 1 have "order a p \<noteq> 0" by (auto split: if_splits)
+  have "order a (pderiv p) = order a e + order a d"
+    using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult)
+  have "order a p = Suc (order a (pderiv p))"
+    using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv)
+  have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> by simp
+  have "([:-a, 1:] ^ (order a (pderiv p))) dvd d"
+    apply (simp add: d)
+    apply (rule dvd_add)
+    apply (rule dvd_mult)
+    apply (simp add: order_divides \<open>p \<noteq> 0\<close>
+           \<open>order a p = Suc (order a (pderiv p))\<close>)
+    apply (rule dvd_mult)
+    apply (simp add: order_divides)
+    done
+  then have "order a (pderiv p) \<le> order a d"
+    using \<open>d \<noteq> 0\<close> by (simp add: order_divides)
+  show ?thesis
+    using \<open>order a p = order a q + order a d\<close>
+    using \<open>order a (pderiv p) = order a e + order a d\<close>
+    using \<open>order a p = Suc (order a (pderiv p))\<close>
+    using \<open>order a (pderiv p) \<le> order a d\<close>
+    by auto
+qed
+
+lemma poly_squarefree_decomp_order2: 
+     "\<lbrakk>pderiv p \<noteq> (0 :: 'a :: field_char_0 poly);
+       p = q * d;
+       pderiv p = e * d;
+       d = r * p + s * pderiv p
+      \<rbrakk> \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+by (blast intro: poly_squarefree_decomp_order)
+
+lemma order_pderiv2: 
+  "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk>
+      \<Longrightarrow> (order a (pderiv p) = n) = (order a p = Suc n)"
+by (auto dest: order_pderiv)
+
+definition rsquarefree :: "'a::idom poly \<Rightarrow> bool"
+  where "rsquarefree p \<longleftrightarrow> p \<noteq> 0 \<and> (\<forall>a. order a p = 0 \<or> order a p = 1)"
+
+lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>h. p = [:h :: 'a :: {semidom,semiring_char_0}:]"
+  by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits)
+
+lemma rsquarefree_roots:
+  fixes p :: "'a :: field_char_0 poly"
+  shows "rsquarefree p = (\<forall>a. \<not>(poly p a = 0 \<and> poly (pderiv p) a = 0))"
+apply (simp add: rsquarefree_def)
+apply (case_tac "p = 0", simp, simp)
+apply (case_tac "pderiv p = 0")
+apply simp
+apply (drule pderiv_iszero, clarsimp)
+apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree)
+apply (force simp add: order_root order_pderiv2)
+  done
+
+lemma poly_squarefree_decomp:
+  assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0"
+    and "p = q * d"
+    and "pderiv p = e * d"
+    and "d = r * p + s * pderiv p"
+  shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
+proof -
+  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
+  with \<open>p = q * d\<close> have "q \<noteq> 0" by simp
+  have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
+    using assms by (rule poly_squarefree_decomp_order2)
+  with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
+    by (simp add: rsquarefree_def order_root)
+qed
+
 
 subsection \<open>Algebraic numbers\<close>
 
@@ -3746,7 +2798,6 @@
   The equivalence is obvious since any rational polynomial can be multiplied with the 
   LCM of its coefficients, yielding an integer polynomial with the same roots.
 \<close>
-subsection \<open>Algebraic numbers\<close>
 
 definition algebraic :: "'a :: field_char_0 \<Rightarrow> bool" where
   "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<int>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
@@ -3761,25 +2812,6 @@
   obtains p where "\<And>i. coeff p i \<in> \<int>" "p \<noteq> 0" "poly p x = 0"
   using assms unfolding algebraic_def by blast
 
-lemma quotient_of_denom_pos': "snd (quotient_of x) > 0"
-  using quotient_of_denom_pos[OF surjective_pairing] .
-  
-lemma of_int_div_in_Ints: 
-  "b dvd a \<Longrightarrow> of_int a div of_int b \<in> (\<int> :: 'a :: ring_div set)"
-proof (cases "of_int b = (0 :: 'a)")
-  assume "b dvd a" "of_int b \<noteq> (0::'a)"
-  then obtain c where "a = b * c" by (elim dvdE)
-  with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
-qed auto
-
-lemma of_int_divide_in_Ints: 
-  "b dvd a \<Longrightarrow> of_int a / of_int b \<in> (\<int> :: 'a :: field set)"
-proof (cases "of_int b = (0 :: 'a)")
-  assume "b dvd a" "of_int b \<noteq> (0::'a)"
-  then obtain c where "a = b * c" by (elim dvdE)
-  with \<open>of_int b \<noteq> (0::'a)\<close> show ?thesis by simp
-qed auto
-
 lemma algebraic_altdef:
   fixes p :: "'a :: field_char_0 poly"
   shows "algebraic x \<longleftrightarrow> (\<exists>p. (\<forall>i. coeff p i \<in> \<rat>) \<and> p \<noteq> 0 \<and> poly p x = 0)"
@@ -3834,285 +2866,1426 @@
 qed
 
 
-text\<open>Lemmas for Derivatives\<close>
-
-lemma order_unique_lemma:
-  fixes p :: "'a::idom poly"
-  assumes "[:-a, 1:] ^ n dvd p" "\<not> [:-a, 1:] ^ Suc n dvd p"
-  shows "n = order a p"
-unfolding Polynomial.order_def
-apply (rule Least_equality [symmetric])
-apply (fact assms)
-apply (rule classical)
-apply (erule notE)
-unfolding not_less_eq_eq
-using assms(1) apply (rule power_le_dvd)
-apply assumption
-done
-
-lemma lemma_order_pderiv1:
-  "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q +
-    smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)"
-apply (simp only: pderiv_mult pderiv_power_Suc)
-apply (simp del: power_Suc of_nat_Suc add: pderiv_pCons)
-done
-
-lemma lemma_order_pderiv:
-  fixes p :: "'a :: field_char_0 poly"
-  assumes n: "0 < n" 
-      and pd: "pderiv p \<noteq> 0" 
-      and pe: "p = [:- a, 1:] ^ n * q" 
-      and nd: "~ [:- a, 1:] dvd q"
-    shows "n = Suc (order a (pderiv p))"
-using n 
+subsection \<open>Content and primitive part of a polynomial\<close>
+
+definition content :: "('a :: semiring_Gcd poly) \<Rightarrow> 'a" where
+  "content p = Gcd (set (coeffs p))"
+
+lemma content_0 [simp]: "content 0 = 0"
+  by (simp add: content_def)
+
+lemma content_1 [simp]: "content 1 = 1"
+  by (simp add: content_def)
+
+lemma content_const [simp]: "content [:c:] = normalize c"
+  by (simp add: content_def cCons_def)
+
+lemma const_poly_dvd_iff_dvd_content:
+  fixes c :: "'a :: semiring_Gcd"
+  shows "[:c:] dvd p \<longleftrightarrow> c dvd content p"
+proof (cases "p = 0")
+  case [simp]: False
+  have "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)" by (rule const_poly_dvd_iff)
+  also have "\<dots> \<longleftrightarrow> (\<forall>a\<in>set (coeffs p). c dvd a)"
+  proof safe
+    fix n :: nat assume "\<forall>a\<in>set (coeffs p). c dvd a"
+    thus "c dvd coeff p n"
+      by (cases "n \<le> degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits)
+  qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits)
+  also have "\<dots> \<longleftrightarrow> c dvd content p"
+    by (simp add: content_def dvd_Gcd_iff mult.commute [of "unit_factor x" for x]
+          dvd_mult_unit_iff)
+  finally show ?thesis .
+qed simp_all
+
+lemma content_dvd [simp]: "[:content p:] dvd p"
+  by (subst const_poly_dvd_iff_dvd_content) simp_all
+  
+lemma content_dvd_coeff [simp]: "content p dvd coeff p n"
+  by (cases "n \<le> degree p") 
+     (auto simp: content_def coeffs_def not_le coeff_eq_0 simp del: upt_Suc intro: Gcd_dvd)
+
+lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
+  by (simp add: content_def Gcd_dvd)
+
+lemma normalize_content [simp]: "normalize (content p) = content p"
+  by (simp add: content_def)
+
+lemma is_unit_content_iff [simp]: "is_unit (content p) \<longleftrightarrow> content p = 1"
+proof
+  assume "is_unit (content p)"
+  hence "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
+  thus "content p = 1" by simp
+qed auto
+
+lemma content_smult [simp]: "content (smult c p) = normalize c * content p"
+  by (simp add: content_def coeffs_smult Gcd_mult)
+
+lemma content_eq_zero_iff [simp]: "content p = 0 \<longleftrightarrow> p = 0"
+  by (auto simp: content_def simp: poly_eq_iff coeffs_def)
+
+definition primitive_part :: "'a :: {semiring_Gcd,idom_divide} poly \<Rightarrow> 'a poly" where
+  "primitive_part p = (if p = 0 then 0 else map_poly (\<lambda>x. x div content p) p)"
+  
+lemma primitive_part_0 [simp]: "primitive_part 0 = 0"
+  by (simp add: primitive_part_def)
+
+lemma content_times_primitive_part [simp]:
+  fixes p :: "'a :: {idom_divide, semiring_Gcd} poly"
+  shows "smult (content p) (primitive_part p) = p"
+proof (cases "p = 0")
+  case False
+  thus ?thesis
+  unfolding primitive_part_def
+  by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs 
+           intro: map_poly_idI)
+qed simp_all
+
+lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
+proof (cases "p = 0")
+  case False
+  hence "primitive_part p = map_poly (\<lambda>x. x div content p) p"
+    by (simp add:  primitive_part_def)
+  also from False have "\<dots> = 0 \<longleftrightarrow> p = 0"
+    by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs)
+  finally show ?thesis using False by simp
+qed simp
+
+lemma content_primitive_part [simp]:
+  assumes "p \<noteq> 0"
+  shows   "content (primitive_part p) = 1"
 proof -
-  have "pderiv ([:- a, 1:] ^ n * q) \<noteq> 0"
-    using assms by auto
-  obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \<noteq> 0"
-    using assms by (cases n) auto
-  have *: "!!k l. k dvd k * pderiv q + smult (of_nat (Suc n')) l \<Longrightarrow> k dvd l"
-    by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff)
-  have "n' = order a (pderiv ([:- a, 1:] ^ Suc n' * q))" 
-  proof (rule order_unique_lemma)
-    show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
-      apply (subst lemma_order_pderiv1)
-      apply (rule dvd_add)
-      apply (metis dvdI dvd_mult2 power_Suc2)
-      apply (metis dvd_smult dvd_triv_right)
-      done
+  have "p = smult (content p) (primitive_part p)" by simp
+  also have "content \<dots> = content p * content (primitive_part p)" 
+    by (simp del: content_times_primitive_part)
+  finally show ?thesis using assms by simp
+qed
+
+lemma content_decompose:
+  fixes p :: "'a :: semiring_Gcd poly"
+  obtains p' where "p = smult (content p) p'" "content p' = 1"
+proof (cases "p = 0")
+  case True
+  thus ?thesis by (intro that[of 1]) simp_all
+next
+  case False
+  from content_dvd[of p] obtain r where r: "p = [:content p:] * r" by (erule dvdE)
+  have "content p * 1 = content p * content r" by (subst r) simp
+  with False have "content r = 1" by (subst (asm) mult_left_cancel) simp_all
+  with r show ?thesis by (intro that[of r]) simp_all
+qed
+
+lemma content_dvd_contentI [intro]:
+  "p dvd q \<Longrightarrow> content p dvd content q"
+  using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast
+  
+lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]"
+  by (simp add: primitive_part_def map_poly_pCons)
+ 
+lemma primitive_part_prim: "content p = 1 \<Longrightarrow> primitive_part p = p"
+  by (auto simp: primitive_part_def)
+  
+lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p"
+proof (cases "p = 0")
+  case False
+  have "p = smult (content p) (primitive_part p)" by simp
+  also from False have "degree \<dots> = degree (primitive_part p)"
+    by (subst degree_smult_eq) simp_all
+  finally show ?thesis ..
+qed simp_all
+
+
+subsection \<open>Division of polynomials\<close>
+
+subsubsection \<open>Division in general\<close>
+  
+instantiation poly :: (idom_divide) idom_divide
+begin
+
+fun divide_poly_main :: "'a \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly 
+  \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly" where
+  "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in 
+     if False \<or> a * lc = cr then (* False \<or> is only because of problem in function-package *)
+     divide_poly_main 
+       lc 
+       (q + mon) 
+       (r - mon * d) 
+       d (dr - 1) n else 0)"
+| "divide_poly_main lc q r d dr 0 = q"
+
+definition divide_poly :: "'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  "divide_poly f g = (if g = 0 then 0 else
+     divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" 
+
+lemma divide_poly_main:
+  assumes d: "d \<noteq> 0" "lc = coeff d (degree d)"
+    and *: "degree (d * r) \<le> dr" "divide_poly_main lc q (d * r) d dr n = q'" 
+    "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> d * r = 0" 
+  shows "q' = q + r"
+  using *
+proof (induct n arbitrary: q r dr)
+  case (Suc n q r dr)
+  let ?rr = "d * r"
+  let ?a = "coeff ?rr dr"
+  let ?qq = "?a div lc"
+  define b where [simp]: "b = monom ?qq n"
+  let ?rrr =  "d * (r - b)"
+  let ?qqq = "q + b"
+  note res = Suc(3)
+  have dr: "dr = n + degree d" using Suc(4) by auto
+  have lc: "lc \<noteq> 0" using d by auto
+  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
+  proof (cases "?qq = 0")
+    case False
+    hence n: "n = degree b" by (simp add: degree_monom_eq)
+    show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum)
+  qed simp
+  also have "\<dots> = lc * coeff b n" unfolding d by simp
+  finally have c2: "coeff (b * d) dr = lc * coeff b n" .
+  have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps)
+  have c1: "coeff (d * r) dr = lc * coeff r n"
+  proof (cases "degree r = n")
+    case True
+    thus ?thesis using Suc(2) unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps)
   next
-    show "\<not> [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)"
-     apply (subst lemma_order_pderiv1)
-     by (metis * nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one)
+    case False
+    have "degree r \<le> n" using dr Suc(2) by auto
+      (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases)
+    with False have r_n: "degree r < n" by auto
+    hence right: "lc * coeff r n = 0" by (simp add: coeff_eq_0)
+    have "coeff (d * r) dr = coeff (d * r) (degree d + n)" unfolding dr by (simp add: ac_simps)
+    also have "\<dots> = 0" using r_n
+      by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 
+        coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq)
+    finally show ?thesis unfolding right .
+  qed
+  have c0: "coeff ?rrr dr = 0" 
+    and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2
+    unfolding b_def coeff_monom coeff_smult c1 using lc by auto
+  from res[unfolded divide_poly_main.simps[of lc q] Let_def] id
+  have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" 
+    by (simp del: divide_poly_main.simps add: field_simps)
+  note IH = Suc(1)[OF _ res] 
+  have dr: "dr = n + degree d" using Suc(4) by auto
+  have deg_rr: "degree ?rr \<le> dr" using Suc(2) by auto
+  have deg_bd: "degree (b * d) \<le> dr"
+    unfolding dr b_def by (rule order.trans[OF degree_mult_le], auto simp: degree_monom_le)
+  have "degree ?rrr \<le> dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd])
+  with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)"
+    by (rule coeff_0_degree_minus_1)
+  have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0"  
+  proof (cases dr)
+    case 0
+    with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto
+    with deg_rrr have "degree ?rrr = 0" by simp
+    from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis
+    show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp
+  qed (insert Suc(4), auto)
+  note IH = IH[OF deg_rrr this]
+  show ?case using IH by simp
+next
+  case (0 q r dr)
+  show ?case 
+  proof (cases "r = 0")
+    case True
+    thus ?thesis using 0 by auto
+  next
+    case False
+    have "degree (d * r) = degree d + degree r" using d False 
+      by (subst degree_mult_eq, auto)
+    thus ?thesis using 0 d by auto
+  qed
+qed 
+
+lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0"
+proof (induct n arbitrary: r d dr)
+  case (Suc n r d dr)
+  show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def
+    by (simp add: Suc del: divide_poly_main.simps)
+qed simp
+
+lemma divide_poly:
+  assumes g: "g \<noteq> 0"
+  shows "(f * g) div g = (f :: 'a poly)" 
+proof - 
+  have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs  g)) 
+    = (f * g) div g" unfolding divide_poly_def Let_def by (simp add: ac_simps)
+  note main = divide_poly_main[OF g refl le_refl this]
+  {
+    fix f :: "'a poly"
+    assume "f \<noteq> 0"
+    hence "length (coeffs f) = Suc (degree f)" unfolding degree_eq_length_coeffs by auto
+  } note len = this
+  have "(f * g) div g = 0 + f"
+  proof (rule main, goal_cases)
+    case 1
+    show ?case
+    proof (cases "f = 0")
+      case True
+      with g show ?thesis by (auto simp: degree_eq_length_coeffs)
+    next
+      case False
+      with g have fg: "g * f \<noteq> 0" by auto
+      show ?thesis unfolding len[OF fg] len[OF g] by auto
+    qed
   qed
-  then show ?thesis
-    by (metis \<open>n = Suc n'\<close> pe)
+  thus ?thesis by simp
+qed
+
+lemma divide_poly_0: "f div 0 = (0 :: 'a poly)"
+  by (simp add: divide_poly_def Let_def divide_poly_main_0)
+
+instance
+  by standard (auto simp: divide_poly divide_poly_0)
+
+end
+
+instance poly :: (idom_divide) algebraic_semidom ..
+
+lemma div_const_poly_conv_map_poly: 
+  assumes "[:c:] dvd p"
+  shows   "p div [:c:] = map_poly (\<lambda>x. x div c) p"
+proof (cases "c = 0")
+  case False
+  from assms obtain q where p: "p = [:c:] * q" by (erule dvdE)
+  moreover {
+    have "smult c q = [:c:] * q" by simp
+    also have "\<dots> div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (insert False, auto)
+    finally have "smult c q div [:c:] = q" .
+  }
+  ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False)
+qed (auto intro!: poly_eqI simp: coeff_map_poly)
+
+lemma is_unit_monom_0:
+  fixes a :: "'a::field"
+  assumes "a \<noteq> 0"
+  shows "is_unit (monom a 0)"
+proof
+  from assms show "1 = monom a 0 * monom (inverse a) 0"
+    by (simp add: mult_monom)
 qed
 
-lemma order_decomp:
-  assumes "p \<noteq> 0"
-  shows "\<exists>q. p = [:- a, 1:] ^ order a p * q \<and> \<not> [:- a, 1:] dvd q"
+lemma is_unit_triv:
+  fixes a :: "'a::field"
+  assumes "a \<noteq> 0"
+  shows "is_unit [:a:]"
+  using assms by (simp add: is_unit_monom_0 monom_0 [symmetric])
+
+lemma is_unit_iff_degree:
+  assumes "p \<noteq> (0 :: _ :: field poly)"
+  shows "is_unit p \<longleftrightarrow> degree p = 0" (is "?P \<longleftrightarrow> ?Q")
+proof
+  assume ?Q
+  then obtain a where "p = [:a:]" by (rule degree_eq_zeroE)
+  with assms show ?P by (simp add: is_unit_triv)
+next
+  assume ?P
+  then obtain q where "q \<noteq> 0" "p * q = 1" ..
+  then have "degree (p * q) = degree 1"
+    by simp
+  with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> have "degree p + degree q = 0"
+    by (simp add: degree_mult_eq)
+  then show ?Q by simp
+qed
+
+lemma is_unit_pCons_iff:
+  "is_unit (pCons (a::_::field) p) \<longleftrightarrow> p = 0 \<and> a \<noteq> 0"
+  by (cases "p = 0") (auto simp add: is_unit_triv is_unit_iff_degree)
+
+lemma is_unit_monom_trival:
+  fixes p :: "'a::field poly"
+  assumes "is_unit p"
+  shows "monom (coeff p (degree p)) 0 = p"
+  using assms by (cases p) (simp_all add: monom_0 is_unit_pCons_iff)
+
+lemma is_unit_const_poly_iff: 
+  "[:c :: 'a :: {comm_semiring_1,semiring_no_zero_divisors}:] dvd 1 \<longleftrightarrow> c dvd 1"
+  by (auto simp: one_poly_def)
+
+lemma is_unit_polyE:
+  fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
+  assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1"
 proof -
-  from assms have A: "[:- a, 1:] ^ order a p dvd p"
-    and B: "\<not> [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order)
-  from A obtain q where C: "p = [:- a, 1:] ^ order a p * q" ..
-  with B have "\<not> [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q"
+  from assms obtain q where "1 = p * q"
+    by (rule dvdE)
+  then have "p \<noteq> 0" and "q \<noteq> 0"
+    by auto
+  from \<open>1 = p * q\<close> have "degree 1 = degree (p * q)"
     by simp
-  then have "\<not> [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q"
-    by simp
-  then have D: "\<not> [:- a, 1:] dvd q"
-    using idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q]
-    by auto
-  from C D show ?thesis by blast
+  also from \<open>p \<noteq> 0\<close> and \<open>q \<noteq> 0\<close> have "\<dots> = degree p + degree q"
+    by (simp add: degree_mult_eq)
+  finally have "degree p = 0" by simp
+  with degree_eq_zeroE obtain c where c: "p = [:c:]" .
+  moreover with \<open>p dvd 1\<close> have "c dvd 1"
+    by (simp add: is_unit_const_poly_iff)
+  ultimately show thesis
+    by (rule that)
+qed
+
+lemma is_unit_polyE':
+  assumes "is_unit (p::_::field poly)"
+  obtains a where "p = monom a 0" and "a \<noteq> 0"
+proof -
+  obtain a q where "p = pCons a q" by (cases p)
+  with assms have "p = [:a:]" and "a \<noteq> 0"
+    by (simp_all add: is_unit_pCons_iff)
+  with that show thesis by (simp add: monom_0)
 qed
 
-lemma order_pderiv:
-  "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk> \<Longrightarrow>
-     (order a p = Suc (order a (pderiv p)))"
-apply (case_tac "p = 0", simp)
-apply (drule_tac a = a and p = p in order_decomp)
-using neq0_conv
-apply (blast intro: lemma_order_pderiv)
-done
-
-lemma order_mult: "p * q \<noteq> 0 \<Longrightarrow> order a (p * q) = order a p + order a q"
+lemma is_unit_poly_iff:
+  fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly"
+  shows "p dvd 1 \<longleftrightarrow> (\<exists>c. p = [:c:] \<and> c dvd 1)"
+  by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff)
+
+  
+subsubsection \<open>Pseudo-Division\<close>
+
+text\<open>This part is by René Thiemann and Akihisa Yamada.\<close>
+
+fun pseudo_divmod_main :: "'a :: comm_ring_1  \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly 
+  \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a poly \<times> 'a poly" where
+  "pseudo_divmod_main lc q r d dr (Suc n) = (let
+     rr = smult lc r;
+     qq = coeff r dr;
+     rrr = rr - monom qq n * d;
+     qqq = smult lc q + monom qq n
+     in pseudo_divmod_main lc qqq rrr d (dr - 1) n)"
+| "pseudo_divmod_main lc q r d dr 0 = (q,r)"
+
+definition pseudo_divmod :: "'a :: comm_ring_1 poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly" where
+  "pseudo_divmod p q \<equiv> if q = 0 then (0,p) else
+     pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))"
+
+lemma pseudo_divmod_main: assumes d: "d \<noteq> 0" "lc = coeff d (degree d)"
+  and *: "degree r \<le> dr" "pseudo_divmod_main lc q r d dr n = (q',r')" 
+    "n = 1 + dr - degree d \<or> dr = 0 \<and> n = 0 \<and> r = 0" 
+  shows "(r' = 0 \<or> degree r' < degree d) \<and> smult (lc^n) (d * q + r) = d * q' + r'"
+  using *
+proof (induct n arbitrary: q r dr)
+  case (Suc n q r dr)
+  let ?rr = "smult lc r"
+  let ?qq = "coeff r dr"
+  define b where [simp]: "b = monom ?qq n"
+  let ?rrr = "?rr - b * d"
+  let ?qqq = "smult lc q + b"
+  note res = Suc(3)
+  from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] 
+  have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" 
+    by (simp del: pseudo_divmod_main.simps)
+  have dr: "dr = n + degree d" using Suc(4) by auto
+  have "coeff (b * d) dr = coeff b n * coeff d (degree d)"
+  proof (cases "?qq = 0")
+    case False
+    hence n: "n = degree b" by (simp add: degree_monom_eq)
+    show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum)
+  qed auto
+  also have "\<dots> = lc * coeff b n" unfolding d by simp
+  finally have "coeff (b * d) dr = lc * coeff b n" .
+  moreover have "coeff ?rr dr = lc * coeff r dr" by simp
+  ultimately have c0: "coeff ?rrr dr = 0" by auto
+  have dr: "dr = n + degree d" using Suc(4) by auto
+  have deg_rr: "degree ?rr \<le> dr" using Suc(2) 
+    using degree_smult_le dual_order.trans by blast 
+  have deg_bd: "degree (b * d) \<le> dr"
+    unfolding dr
+    by(rule order.trans[OF degree_mult_le], auto simp: degree_monom_le)
+  have "degree ?rrr \<le> dr"
+    using degree_diff_le[OF deg_rr deg_bd] by auto
+  with c0 have deg_rrr: "degree ?rrr \<le> (dr - 1)" by (rule coeff_0_degree_minus_1)
+  have "n = 1 + (dr - 1) - degree d \<or> dr - 1 = 0 \<and> n = 0 \<and> ?rrr = 0"
+  proof (cases dr)
+    case 0
+    with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto
+    with deg_rrr have "degree ?rrr = 0" by simp
+    hence "\<exists> a. ?rrr = [: a :]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases)
+    from this obtain a where rrr: "?rrr = [:a:]" by auto
+    show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp
+  qed (insert Suc(4), auto)
+  note IH = Suc(1)[OF deg_rrr res this]
+  show ?case
+  proof (intro conjI)
+    show "r' = 0 \<or> degree r' < degree d" using IH by blast
+    show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'"
+      unfolding IH[THEN conjunct2,symmetric]
+      by (simp add: field_simps smult_add_right)
+  qed
+qed auto
+
+lemma pseudo_divmod:
+  assumes g: "g \<noteq> 0" and *: "pseudo_divmod f g = (q,r)" 
+  shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A)
+    and "r = 0 \<or> degree r < degree g" (is ?B)
 proof -
-  define i where "i = order a p"
-  define j where "j = order a q"
-  define t where "t = [:-a, 1:]"
-  have t_dvd_iff: "\<And>u. t dvd u \<longleftrightarrow> poly u a = 0"
-    unfolding t_def by (simp add: dvd_iff_poly_eq_0)
-  assume "p * q \<noteq> 0"
-  then show "order a (p * q) = i + j"
-    apply clarsimp
-    apply (drule order [where a=a and p=p, folded i_def t_def])
-    apply (drule order [where a=a and p=q, folded j_def t_def])
-    apply clarify
-    apply (erule dvdE)+
-    apply (rule order_unique_lemma [symmetric], fold t_def)
-    apply (simp_all add: power_add t_dvd_iff)
-    done
+  from *[unfolded pseudo_divmod_def Let_def]
+  have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g)
+  note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl]
+  have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \<or>
+    degree f = 0 \<and> 1 + length (coeffs f) - length (coeffs g) = 0 \<and> f = 0" using g 
+    by (cases "f = 0"; cases "coeffs g", auto simp: degree_eq_length_coeffs)
+  note main = main[OF this]
+  from main show "r = 0 \<or> degree r < degree g" by auto
+  show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" 
+    by (subst main[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs,
+    insert g, cases "f = 0"; cases "coeffs g", auto)
+qed
+  
+definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)"
+
+lemma snd_pseudo_divmod_main:
+  "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)"
+by (induct n arbitrary: q q' lc r d dr; simp add: Let_def)
+
+definition pseudo_mod 
+    :: "'a :: {comm_ring_1,semiring_1_no_zero_divisors} poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  "pseudo_mod f g = snd (pseudo_divmod f g)"
+  
+lemma pseudo_mod:
+  fixes f g
+  defines "r \<equiv> pseudo_mod f g"
+  assumes g: "g \<noteq> 0"
+  shows "\<exists> a q. a \<noteq> 0 \<and> smult a f = g * q + r" "r = 0 \<or> degree r < degree g"
+proof - 
+  let ?cg = "coeff g (degree g)"
+  let ?cge = "?cg ^ (Suc (degree f) - degree g)"
+  define a where "a = ?cge"
+  obtain q where pdm: "pseudo_divmod f g = (q,r)" using r_def[unfolded pseudo_mod_def]
+    by (cases "pseudo_divmod f g", auto)
+  from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \<or> degree r < degree g" 
+    unfolding a_def by auto
+  show "r = 0 \<or> degree r < degree g" by fact
+  from g have "a \<noteq> 0" unfolding a_def by auto
+  thus "\<exists> a q. a \<noteq> 0 \<and> smult a f = g * q + r" using id by auto
+qed
+  
+lemma fst_pseudo_divmod_main_as_divide_poly_main:
+  assumes d: "d \<noteq> 0"
+  defines lc: "lc \<equiv> coeff d (degree d)"
+  shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n"
+proof(induct n arbitrary: q r dr)
+  case 0 then show ?case by simp
+next
+  case (Suc n)
+    note lc0 = leading_coeff_neq_0[OF d, folded lc]
+    then have "pseudo_divmod_main lc q r d dr (Suc n) =
+    pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n)
+      (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n"
+    by (simp add: Let_def ac_simps)
+    also have "fst ... = divide_poly_main lc
+     (smult (lc^n) (smult lc q + monom (coeff r dr) n))
+     (smult (lc^n) (smult lc r - monom (coeff r dr) n * d))
+     d (dr - 1) n"
+      unfolding Suc[unfolded divide_poly_main.simps Let_def]..
+    also have "... = divide_poly_main lc (smult (lc ^ Suc n) q)
+        (smult (lc ^ Suc n) r) d dr (Suc n)"
+      unfolding smult_monom smult_distribs mult_smult_left[symmetric]
+      using lc0 by (simp add: Let_def ac_simps)
+    finally show ?case.
 qed
 
-lemma order_smult:
-  assumes "c \<noteq> 0" 
-  shows "order x (smult c p) = order x p"
+
+subsubsection \<open>Division in polynomials over fields\<close>
+
+lemma pseudo_divmod_field:
+  assumes g: "(g::'a::field poly) \<noteq> 0" and *: "pseudo_divmod f g = (q,r)"
+  defines "c \<equiv> coeff g (degree g) ^ (Suc (degree f) - degree g)"
+  shows "f = g * smult (1/c) q + smult (1/c) r"
+proof -
+  from leading_coeff_neq_0[OF g] have c0: "c \<noteq> 0" unfolding c_def by auto
+  from pseudo_divmod(1)[OF g *, folded c_def]
+  have "smult c f = g * q + r" by auto
+  also have "smult (1/c) ... = g * smult (1/c) q + smult (1/c) r" by (simp add: smult_add_right)
+  finally show ?thesis using c0 by auto
+qed
+
+lemma divide_poly_main_field:
+  assumes d: "(d::'a::field poly) \<noteq> 0"
+  defines lc: "lc \<equiv> coeff d (degree d)"
+  shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1/lc)^n) q) (smult ((1/lc)^n) r) d dr n)"
+  unfolding lc
+  by(subst fst_pseudo_divmod_main_as_divide_poly_main, auto simp: d power_one_over)
+
+lemma divide_poly_field:
+  fixes f g :: "'a :: field poly"
+  defines "f' \<equiv> smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f"
+  shows "(f::'a::field poly) div g = fst (pseudo_divmod f' g)"
+proof (cases "g = 0")
+  case True show ?thesis 
+    unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0)
+next
+  case False
+    from leading_coeff_neq_0[OF False] have "degree f' = degree f" unfolding f'_def by auto
+    then show ?thesis
+      using length_coeffs_degree[of f'] length_coeffs_degree[of f]
+      unfolding divide_poly_def pseudo_divmod_def Let_def
+                divide_poly_main_field[OF False]
+                length_coeffs_degree[OF False] 
+                f'_def
+      by force
+qed
+
+instantiation poly :: ("{normalization_semidom, idom_divide}") normalization_semidom
+begin
+
+definition unit_factor_poly :: "'a poly \<Rightarrow> 'a poly"
+  where "unit_factor_poly p = monom (unit_factor (lead_coeff p)) 0"
+
+definition normalize_poly :: "'a poly \<Rightarrow> 'a poly"
+  where "normalize_poly p = map_poly (\<lambda>x. x div unit_factor (lead_coeff p)) p"
+
+instance proof
+  fix p :: "'a poly"
+  show "unit_factor p * normalize p = p"
+    by (cases "p = 0")
+       (simp_all add: unit_factor_poly_def normalize_poly_def monom_0 
+          smult_conv_map_poly map_poly_map_poly o_def)
+next
+  fix p :: "'a poly"
+  assume "is_unit p"
+  then obtain c where p: "p = [:c:]" "is_unit c"
+    by (auto simp: is_unit_poly_iff)
+  thus "normalize p = 1"
+    by (simp add: normalize_poly_def map_poly_pCons is_unit_normalize one_poly_def)
+next
+  fix p :: "'a poly" assume "p \<noteq> 0"
+  thus "is_unit (unit_factor p)"
+    by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff)
+qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult)
+
+end
+
+lemma normalize_poly_eq_div:
+  "normalize p = p div [:unit_factor (lead_coeff p):]"
 proof (cases "p = 0")
   case False
+  thus ?thesis
+    by (subst div_const_poly_conv_map_poly)
+       (auto simp: normalize_poly_def const_poly_dvd_iff)
+qed (auto simp: normalize_poly_def)
+
+lemma unit_factor_pCons:
+  "unit_factor (pCons a p) = (if p = 0 then monom (unit_factor a) 0 else unit_factor p)"
+  by (simp add: unit_factor_poly_def)
+
+lemma normalize_monom [simp]:
+  "normalize (monom a n) = monom (normalize a) n"
+  by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_def degree_monom_eq)
+
+lemma unit_factor_monom [simp]:
+  "unit_factor (monom a n) = monom (unit_factor a) 0"
+  by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq)
+
+lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]"
+  by (simp add: normalize_poly_def map_poly_pCons)
+
+lemma normalize_smult: "normalize (smult c p) = smult (normalize c) (normalize p)"
+proof -
   have "smult c p = [:c:] * p" by simp
-  also from assms False have "order x \<dots> = order x [:c:] + order x p" 
-    by (subst order_mult) simp_all
-  also from assms have "order x [:c:] = 0" by (intro order_0I) auto
-  finally show ?thesis by simp
-qed simp
-
-(* Next two lemmas contributed by Wenda Li *)
-lemma order_1_eq_0 [simp]:"order x 1 = 0" 
-  by (metis order_root poly_1 zero_neq_one)
-
-lemma order_power_n_n: "order a ([:-a,1:]^n)=n" 
-proof (induct n) (*might be proved more concisely using nat_less_induct*)
-  case 0
-  thus ?case by (metis order_root poly_1 power_0 zero_neq_one)
-next 
-  case (Suc n)
-  have "order a ([:- a, 1:] ^ Suc n)=order a ([:- a, 1:] ^ n) + order a [:-a,1:]" 
-    by (metis (no_types, hide_lams) One_nat_def add_Suc_right monoid_add_class.add.right_neutral 
-      one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right)
-  moreover have "order a [:-a,1:]=1" unfolding order_def
-    proof (rule Least_equality,rule ccontr)
-      assume  "\<not> \<not> [:- a, 1:] ^ Suc 1 dvd [:- a, 1:]"
-      hence "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" by simp
-      hence "degree ([:- a, 1:] ^ Suc 1) \<le> degree ([:- a, 1:] )" 
-        by (rule dvd_imp_degree_le,auto) 
-      thus False by auto
-    next
-      fix y assume asm:"\<not> [:- a, 1:] ^ Suc y dvd [:- a, 1:]"
-      show "1 \<le> y" 
-        proof (rule ccontr)
-          assume "\<not> 1 \<le> y"
-          hence "y=0" by auto
-          hence "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto
-          thus False using asm by auto
-        qed
-    qed
-  ultimately show ?case using Suc by auto
+  also have "normalize \<dots> = smult (normalize c) (normalize p)"
+    by (subst normalize_mult) (simp add: normalize_const_poly)
+  finally show ?thesis .
 qed
 
-lemma order_0_monom [simp]:
-  assumes "c \<noteq> 0"
-  shows   "order 0 (monom c n) = n"
-  using assms order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult)
-
-lemma dvd_imp_order_le:
-  "q \<noteq> 0 \<Longrightarrow> p dvd q \<Longrightarrow> Polynomial.order a p \<le> Polynomial.order a q"
-  by (auto simp: order_mult elim: dvdE)
-
-text\<open>Now justify the standard squarefree decomposition, i.e. f / gcd(f,f').\<close>
-
-lemma order_divides: "[:-a, 1:] ^ n dvd p \<longleftrightarrow> p = 0 \<or> n \<le> order a p"
-apply (cases "p = 0", auto)
-apply (drule order_2 [where a=a and p=p])
-apply (metis not_less_eq_eq power_le_dvd)
-apply (erule power_le_dvd [OF order_1])
-done
-
-lemma monom_1_dvd_iff:
-  assumes "p \<noteq> 0"
-  shows   "monom 1 n dvd p \<longleftrightarrow> n \<le> Polynomial.order 0 p"
-  using assms order_divides[of 0 n p] by (simp add: monom_altdef)
-
-lemma poly_squarefree_decomp_order:
-  assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0"
-  and p: "p = q * d"
-  and p': "pderiv p = e * d"
-  and d: "d = r * p + s * pderiv p"
-  shows "order a q = (if order a p = 0 then 0 else 1)"
-proof (rule classical)
-  assume 1: "order a q \<noteq> (if order a p = 0 then 0 else 1)"
-  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
-  with p have "order a p = order a q + order a d"
-    by (simp add: order_mult)
-  with 1 have "order a p \<noteq> 0" by (auto split: if_splits)
-  have "order a (pderiv p) = order a e + order a d"
-    using \<open>pderiv p \<noteq> 0\<close> \<open>pderiv p = e * d\<close> by (simp add: order_mult)
-  have "order a p = Suc (order a (pderiv p))"
-    using \<open>pderiv p \<noteq> 0\<close> \<open>order a p \<noteq> 0\<close> by (rule order_pderiv)
-  have "d \<noteq> 0" using \<open>p \<noteq> 0\<close> \<open>p = q * d\<close> by simp
-  have "([:-a, 1:] ^ (order a (pderiv p))) dvd d"
-    apply (simp add: d)
-    apply (rule dvd_add)
-    apply (rule dvd_mult)
-    apply (simp add: order_divides \<open>p \<noteq> 0\<close>
-           \<open>order a p = Suc (order a (pderiv p))\<close>)
-    apply (rule dvd_mult)
-    apply (simp add: order_divides)
-    done
-  then have "order a (pderiv p) \<le> order a d"
-    using \<open>d \<noteq> 0\<close> by (simp add: order_divides)
-  show ?thesis
-    using \<open>order a p = order a q + order a d\<close>
-    using \<open>order a (pderiv p) = order a e + order a d\<close>
-    using \<open>order a p = Suc (order a (pderiv p))\<close>
-    using \<open>order a (pderiv p) \<le> order a d\<close>
-    by auto
-qed
-
-lemma poly_squarefree_decomp_order2: 
-     "\<lbrakk>pderiv p \<noteq> (0 :: 'a :: field_char_0 poly);
-       p = q * d;
-       pderiv p = e * d;
-       d = r * p + s * pderiv p
-      \<rbrakk> \<Longrightarrow> \<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-by (blast intro: poly_squarefree_decomp_order)
-
-lemma order_pderiv2: 
-  "\<lbrakk>pderiv p \<noteq> 0; order a (p :: 'a :: field_char_0 poly) \<noteq> 0\<rbrakk>
-      \<Longrightarrow> (order a (pderiv p) = n) = (order a p = Suc n)"
-by (auto dest: order_pderiv)
-
-definition
-  rsquarefree :: "'a::idom poly => bool" where
-  "rsquarefree p = (p \<noteq> 0 & (\<forall>a. (order a p = 0) | (order a p = 1)))"
-
-lemma pderiv_iszero: "pderiv p = 0 \<Longrightarrow> \<exists>h. p = [:h :: 'a :: {semidom,semiring_char_0}:]"
-  by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits)
-
-lemma rsquarefree_roots:
-  fixes p :: "'a :: field_char_0 poly"
-  shows "rsquarefree p = (\<forall>a. \<not>(poly p a = 0 \<and> poly (pderiv p) a = 0))"
-apply (simp add: rsquarefree_def)
-apply (case_tac "p = 0", simp, simp)
-apply (case_tac "pderiv p = 0")
-apply simp
-apply (drule pderiv_iszero, clarsimp)
-apply (metis coeff_0 coeff_pCons_0 degree_pCons_0 le0 le_antisym order_degree)
-apply (force simp add: order_root order_pderiv2)
-done
-
-lemma poly_squarefree_decomp:
-  assumes "pderiv (p :: 'a :: field_char_0 poly) \<noteq> 0"
-    and "p = q * d"
-    and "pderiv p = e * d"
-    and "d = r * p + s * pderiv p"
-  shows "rsquarefree q & (\<forall>a. (poly q a = 0) = (poly p a = 0))"
+lemma smult_content_normalize_primitive_part [simp]:
+  "smult (content p) (normalize (primitive_part p)) = normalize p"
 proof -
-  from \<open>pderiv p \<noteq> 0\<close> have "p \<noteq> 0" by auto
-  with \<open>p = q * d\<close> have "q \<noteq> 0" by simp
-  have "\<forall>a. order a q = (if order a p = 0 then 0 else 1)"
-    using assms by (rule poly_squarefree_decomp_order2)
-  with \<open>p \<noteq> 0\<close> \<open>q \<noteq> 0\<close> show ?thesis
-    by (simp add: rsquarefree_def order_root)
-qed
-
-lemma coeff_monom_mult: 
-  "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))"
-proof -
-  have "coeff (monom c n * p) k = (\<Sum>i\<le>k. (if n = i then c else 0) * coeff p (k - i))"
-    by (simp add: coeff_mult)
-  also have "\<dots> = (\<Sum>i\<le>k. (if n = i then c * coeff p (k - i) else 0))"
-    by (intro sum.cong) simp_all
-  also have "\<dots> = (if k < n then 0 else c * coeff p (k - n))" by (simp add: sum.delta')
+  have "smult (content p) (normalize (primitive_part p)) = 
+          normalize ([:content p:] * primitive_part p)" 
+    by (subst normalize_mult) (simp_all add: normalize_const_poly)
+  also have "[:content p:] * primitive_part p = p" by simp
   finally show ?thesis .
 qed
 
-lemma monom_1_dvd_iff':
-  "monom 1 n dvd p \<longleftrightarrow> (\<forall>k<n. coeff p k = 0)"
+inductive eucl_rel_poly :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly \<Rightarrow> bool"
+  where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)"
+  | eucl_rel_poly_dividesI: "y \<noteq> 0 \<Longrightarrow> x = q * y \<Longrightarrow> eucl_rel_poly x y (q, 0)"
+  | eucl_rel_poly_remainderI: "y \<noteq> 0 \<Longrightarrow> degree r < degree y
+      \<Longrightarrow> x = q * y + r \<Longrightarrow> eucl_rel_poly x y (q, r)"
+  
+lemma eucl_rel_poly_iff:
+  "eucl_rel_poly x y (q, r) \<longleftrightarrow>
+    x = q * y + r \<and>
+      (if y = 0 then q = 0 else r = 0 \<or> degree r < degree y)"
+  by (auto elim: eucl_rel_poly.cases
+    intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI)
+  
+lemma eucl_rel_poly_0:
+  "eucl_rel_poly 0 y (0, 0)"
+  unfolding eucl_rel_poly_iff by simp
+
+lemma eucl_rel_poly_by_0:
+  "eucl_rel_poly x 0 (0, x)"
+  unfolding eucl_rel_poly_iff by simp
+
+lemma eucl_rel_poly_pCons:
+  assumes rel: "eucl_rel_poly x y (q, r)"
+  assumes y: "y \<noteq> 0"
+  assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)"
+  shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)"
+    (is "eucl_rel_poly ?x y (?q, ?r)")
+proof -
+  have x: "x = q * y + r" and r: "r = 0 \<or> degree r < degree y"
+    using assms unfolding eucl_rel_poly_iff by simp_all
+
+  have 1: "?x = ?q * y + ?r"
+    using b x by simp
+
+  have 2: "?r = 0 \<or> degree ?r < degree y"
+  proof (rule eq_zero_or_degree_less)
+    show "degree ?r \<le> degree y"
+    proof (rule degree_diff_le)
+      show "degree (pCons a r) \<le> degree y"
+        using r by auto
+      show "degree (smult b y) \<le> degree y"
+        by (rule degree_smult_le)
+    qed
+  next
+    show "coeff ?r (degree y) = 0"
+      using \<open>y \<noteq> 0\<close> unfolding b by simp
+  qed
+
+  from 1 2 show ?thesis
+    unfolding eucl_rel_poly_iff
+    using \<open>y \<noteq> 0\<close> by simp
+qed
+
+lemma eucl_rel_poly_exists: "\<exists>q r. eucl_rel_poly x y (q, r)"
+apply (cases "y = 0")
+apply (fast intro!: eucl_rel_poly_by_0)
+apply (induct x)
+apply (fast intro!: eucl_rel_poly_0)
+apply (fast intro!: eucl_rel_poly_pCons)
+done
+
+lemma eucl_rel_poly_unique:
+  assumes 1: "eucl_rel_poly x y (q1, r1)"
+  assumes 2: "eucl_rel_poly x y (q2, r2)"
+  shows "q1 = q2 \<and> r1 = r2"
+proof (cases "y = 0")
+  assume "y = 0" with assms show ?thesis
+    by (simp add: eucl_rel_poly_iff)
+next
+  assume [simp]: "y \<noteq> 0"
+  from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \<or> degree r1 < degree y"
+    unfolding eucl_rel_poly_iff by simp_all
+  from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \<or> degree r2 < degree y"
+    unfolding eucl_rel_poly_iff by simp_all
+  from q1 q2 have q3: "(q1 - q2) * y = r2 - r1"
+    by (simp add: algebra_simps)
+  from r1 r2 have r3: "(r2 - r1) = 0 \<or> degree (r2 - r1) < degree y"
+    by (auto intro: degree_diff_less)
+
+  show "q1 = q2 \<and> r1 = r2"
+  proof (rule ccontr)
+    assume "\<not> (q1 = q2 \<and> r1 = r2)"
+    with q3 have "q1 \<noteq> q2" and "r1 \<noteq> r2" by auto
+    with r3 have "degree (r2 - r1) < degree y" by simp
+    also have "degree y \<le> degree (q1 - q2) + degree y" by simp
+    also have "\<dots> = degree ((q1 - q2) * y)"
+      using \<open>q1 \<noteq> q2\<close> by (simp add: degree_mult_eq)
+    also have "\<dots> = degree (r2 - r1)"
+      using q3 by simp
+    finally have "degree (r2 - r1) < degree (r2 - r1)" .
+    then show "False" by simp
+  qed
+qed
+
+lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \<longleftrightarrow> q = 0 \<and> r = 0"
+by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0)
+
+lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \<longleftrightarrow> q = 0 \<and> r = x"
+by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0)
+
+lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1]
+
+lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2]
+
+instantiation poly :: (field) ring_div
+begin
+
+definition modulo_poly where
+  mod_poly_def: "f mod g \<equiv>
+    if g = 0 then f
+    else pseudo_mod (smult ((1/coeff g (degree g)) ^ (Suc (degree f) - degree g)) f) g"
+
+lemma eucl_rel_poly: "eucl_rel_poly (x::'a::field poly) y (x div y, x mod y)"
+  unfolding eucl_rel_poly_iff
+proof (intro conjI)
+  show "x = x div y * y + x mod y"
+  proof(cases "y = 0")
+    case True show ?thesis by(simp add: True divide_poly_def divide_poly_0 mod_poly_def)
+  next
+    case False
+    then have "pseudo_divmod (smult ((1 / coeff y (degree y)) ^ (Suc (degree x) - degree y)) x) y =
+          (x div y, x mod y)"
+      unfolding divide_poly_field mod_poly_def pseudo_mod_def by simp
+    from pseudo_divmod[OF False this]
+    show ?thesis using False
+      by (simp add: power_mult_distrib[symmetric] mult.commute)
+  qed
+  show "if y = 0 then x div y = 0 else x mod y = 0 \<or> degree (x mod y) < degree y"
+  proof (cases "y = 0")
+    case True then show ?thesis by auto
+  next
+    case False
+      with pseudo_mod[OF this] show ?thesis unfolding mod_poly_def by simp
+  qed
+qed
+
+lemma div_poly_eq:
+  "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x div y = q"
+  by(rule eucl_rel_poly_unique_div[OF eucl_rel_poly])
+
+lemma mod_poly_eq:
+  "eucl_rel_poly (x::'a::field poly) y (q, r) \<Longrightarrow> x mod y = r"
+  by (rule eucl_rel_poly_unique_mod[OF eucl_rel_poly])
+
+instance
 proof
-  assume "monom 1 n dvd p"
-  then obtain r where r: "p = monom 1 n * r" by (elim dvdE)
-  thus "\<forall>k<n. coeff p k = 0" by (simp add: coeff_mult)
+  fix x y :: "'a poly"
+  from eucl_rel_poly[of x y,unfolded eucl_rel_poly_iff]
+  show "x div y * y + x mod y = x" by auto
+next
+  fix x y z :: "'a poly"
+  assume "y \<noteq> 0"
+  hence "eucl_rel_poly (x + z * y) y (z + x div y, x mod y)"
+    using eucl_rel_poly [of x y]
+    by (simp add: eucl_rel_poly_iff distrib_right)
+  thus "(x + z * y) div y = z + x div y"
+    by (rule div_poly_eq)
 next
-  assume zero: "(\<forall>k<n. coeff p k = 0)"
-  define r where "r = Abs_poly (\<lambda>k. coeff p (k + n))"
-  have "\<forall>\<^sub>\<infinity>k. coeff p (k + n) = 0"
-    by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, 
-        subst cofinite_eq_sequentially [symmetric]) transfer
-  hence coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def
-    by (subst poly.Abs_poly_inverse) simp_all
-  have "p = monom 1 n * r"
-    by (intro poly_eqI, subst coeff_monom_mult) (insert zero, simp_all)
-  thus "monom 1 n dvd p" by simp
+  fix x y z :: "'a poly"
+  assume "x \<noteq> 0"
+  show "(x * y) div (x * z) = y div z"
+  proof (cases "y \<noteq> 0 \<and> z \<noteq> 0")
+    have "\<And>x::'a poly. eucl_rel_poly x 0 (0, x)"
+      by (rule eucl_rel_poly_by_0)
+    then have [simp]: "\<And>x::'a poly. x div 0 = 0"
+      by (rule div_poly_eq)
+    have "\<And>x::'a poly. eucl_rel_poly 0 x (0, 0)"
+      by (rule eucl_rel_poly_0)
+    then have [simp]: "\<And>x::'a poly. 0 div x = 0"
+      by (rule div_poly_eq)
+    case False then show ?thesis by auto
+  next
+    case True then have "y \<noteq> 0" and "z \<noteq> 0" by auto
+    with \<open>x \<noteq> 0\<close>
+    have "\<And>q r. eucl_rel_poly y z (q, r) \<Longrightarrow> eucl_rel_poly (x * y) (x * z) (q, x * r)"
+      by (auto simp add: eucl_rel_poly_iff algebra_simps)
+        (rule classical, simp add: degree_mult_eq)
+    moreover from eucl_rel_poly have "eucl_rel_poly y z (y div z, y mod z)" .
+    ultimately have "eucl_rel_poly (x * y) (x * z) (y div z, x * (y mod z))" .
+    then show ?thesis by (simp add: div_poly_eq)
+  qed
+qed
+
+end
+
+lemma degree_mod_less:
+  "y \<noteq> 0 \<Longrightarrow> x mod y = 0 \<or> degree (x mod y) < degree y"
+  using eucl_rel_poly [of x y]
+  unfolding eucl_rel_poly_iff by simp
+
+lemma degree_mod_less': "b \<noteq> 0 \<Longrightarrow> a mod b \<noteq> 0 \<Longrightarrow> degree (a mod b) < degree b"
+  using degree_mod_less[of b a] by auto
+
+lemma div_poly_less: "degree (x::'a::field poly) < degree y \<Longrightarrow> x div y = 0"
+proof -
+  assume "degree x < degree y"
+  hence "eucl_rel_poly x y (0, x)"
+    by (simp add: eucl_rel_poly_iff)
+  thus "x div y = 0" by (rule div_poly_eq)
+qed
+
+lemma mod_poly_less: "degree x < degree y \<Longrightarrow> x mod y = x"
+proof -
+  assume "degree x < degree y"
+  hence "eucl_rel_poly x y (0, x)"
+    by (simp add: eucl_rel_poly_iff)
+  thus "x mod y = x" by (rule mod_poly_eq)
+qed
+
+lemma eucl_rel_poly_smult_left:
+  "eucl_rel_poly x y (q, r)
+    \<Longrightarrow> eucl_rel_poly (smult a x) y (smult a q, smult a r)"
+  unfolding eucl_rel_poly_iff by (simp add: smult_add_right)
+
+lemma div_smult_left: "(smult (a::'a::field) x) div y = smult a (x div y)"
+  by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
+
+lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)"
+  by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly)
+
+lemma poly_div_minus_left [simp]:
+  fixes x y :: "'a::field poly"
+  shows "(- x) div y = - (x div y)"
+  using div_smult_left [of "- 1::'a"] by simp
+
+lemma poly_mod_minus_left [simp]:
+  fixes x y :: "'a::field poly"
+  shows "(- x) mod y = - (x mod y)"
+  using mod_smult_left [of "- 1::'a"] by simp
+
+lemma eucl_rel_poly_add_left:
+  assumes "eucl_rel_poly x y (q, r)"
+  assumes "eucl_rel_poly x' y (q', r')"
+  shows "eucl_rel_poly (x + x') y (q + q', r + r')"
+  using assms unfolding eucl_rel_poly_iff
+  by (auto simp add: algebra_simps degree_add_less)
+
+lemma poly_div_add_left:
+  fixes x y z :: "'a::field poly"
+  shows "(x + y) div z = x div z + y div z"
+  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
+  by (rule div_poly_eq)
+
+lemma poly_mod_add_left:
+  fixes x y z :: "'a::field poly"
+  shows "(x + y) mod z = x mod z + y mod z"
+  using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly]
+  by (rule mod_poly_eq)
+
+lemma poly_div_diff_left:
+  fixes x y z :: "'a::field poly"
+  shows "(x - y) div z = x div z - y div z"
+  by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left)
+
+lemma poly_mod_diff_left:
+  fixes x y z :: "'a::field poly"
+  shows "(x - y) mod z = x mod z - y mod z"
+  by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left)
+
+lemma eucl_rel_poly_smult_right:
+  "a \<noteq> 0 \<Longrightarrow> eucl_rel_poly x y (q, r)
+    \<Longrightarrow> eucl_rel_poly x (smult a y) (smult (inverse a) q, r)"
+  unfolding eucl_rel_poly_iff by simp
+
+lemma div_smult_right:
+  "(a::'a::field) \<noteq> 0 \<Longrightarrow> x div (smult a y) = smult (inverse a) (x div y)"
+  by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
+
+lemma mod_smult_right: "a \<noteq> 0 \<Longrightarrow> x mod (smult a y) = x mod y"
+  by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly)
+
+lemma poly_div_minus_right [simp]:
+  fixes x y :: "'a::field poly"
+  shows "x div (- y) = - (x div y)"
+  using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq)
+
+lemma poly_mod_minus_right [simp]:
+  fixes x y :: "'a::field poly"
+  shows "x mod (- y) = x mod y"
+  using mod_smult_right [of "- 1::'a"] by simp
+
+lemma eucl_rel_poly_mult:
+  "eucl_rel_poly x y (q, r) \<Longrightarrow> eucl_rel_poly q z (q', r')
+    \<Longrightarrow> eucl_rel_poly x (y * z) (q', y * r' + r)"
+apply (cases "z = 0", simp add: eucl_rel_poly_iff)
+apply (cases "y = 0", simp add: eucl_rel_poly_by_0_iff eucl_rel_poly_0_iff)
+apply (cases "r = 0")
+apply (cases "r' = 0")
+apply (simp add: eucl_rel_poly_iff)
+apply (simp add: eucl_rel_poly_iff field_simps degree_mult_eq)
+apply (cases "r' = 0")
+apply (simp add: eucl_rel_poly_iff degree_mult_eq)
+apply (simp add: eucl_rel_poly_iff field_simps)
+apply (simp add: degree_mult_eq degree_add_less)
+done
+
+lemma poly_div_mult_right:
+  fixes x y z :: "'a::field poly"
+  shows "x div (y * z) = (x div y) div z"
+  by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
+
+lemma poly_mod_mult_right:
+  fixes x y z :: "'a::field poly"
+  shows "x mod (y * z) = y * (x div y mod z) + x mod y"
+  by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+)
+
+lemma mod_pCons:
+  fixes a and x
+  assumes y: "y \<noteq> 0"
+  defines b: "b \<equiv> coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)"
+  shows "(pCons a x) mod y = (pCons a (x mod y) - smult b y)"
+unfolding b
+apply (rule mod_poly_eq)
+apply (rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl])
+done
+
+definition pdivmod :: "'a::field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly \<times> 'a poly"
+where
+  "pdivmod p q = (p div q, p mod q)"
+
+lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \<longleftrightarrow> pdivmod p q = (r, s)"
+  by (metis pdivmod_def eucl_rel_poly eucl_rel_poly_unique)
+
+lemma pdivmod_0:
+  "pdivmod 0 q = (0, 0)"
+  by (simp add: pdivmod_def)
+
+lemma pdivmod_pCons:
+  "pdivmod (pCons a p) q =
+    (if q = 0 then (0, pCons a p) else
+      (let (s, r) = pdivmod p q;
+           b = coeff (pCons a r) (degree q) / coeff q (degree q)
+        in (pCons b s, pCons a r - smult b q)))"
+  apply (simp add: pdivmod_def Let_def, safe)
+  apply (rule div_poly_eq)
+  apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
+  apply (rule mod_poly_eq)
+  apply (erule eucl_rel_poly_pCons [OF eucl_rel_poly _ refl])
+  done
+
+lemma pdivmod_fold_coeffs:
+  "pdivmod p q = (if q = 0 then (0, p)
+    else fold_coeffs (\<lambda>a (s, r).
+      let b = coeff (pCons a r) (degree q) / coeff q (degree q)
+      in (pCons b s, pCons a r - smult b q)
+   ) p (0, 0))"
+  apply (cases "q = 0")
+  apply (simp add: pdivmod_def)
+  apply (rule sym)
+  apply (induct p)
+  apply (simp_all add: pdivmod_0 pdivmod_pCons)
+  apply (case_tac "a = 0 \<and> p = 0")
+  apply (auto simp add: pdivmod_def)
+  done
+
+    
+subsubsection \<open>List-based versions for fast implementation\<close>
+(* Subsection by:
+      Sebastiaan Joosten
+      René Thiemann
+      Akihisa Yamada
+    *)
+fun minus_poly_rev_list :: "'a :: group_add list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)"
+| "minus_poly_rev_list xs [] = xs"
+| "minus_poly_rev_list [] (y # ys) = []"
+
+fun pseudo_divmod_main_list :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list 
+  \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" where
+  "pseudo_divmod_main_list lc q r d (Suc n) = (let
+     rr = map (op * lc) r;
+     a = hd r;
+     qqq = cCons a (map (op * lc) q);
+     rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d))
+     in pseudo_divmod_main_list lc qqq rrr d n)"
+| "pseudo_divmod_main_list lc q r d 0 = (q,r)"
+
+fun pseudo_mod_main_list :: "'a::comm_ring_1 \<Rightarrow> 'a list \<Rightarrow> 'a list 
+  \<Rightarrow> nat \<Rightarrow> 'a list" where
+  "pseudo_mod_main_list lc r d (Suc n) = (let
+     rr = map (op * lc) r;
+     a = hd r;
+     rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map (op * a) d))
+     in pseudo_mod_main_list lc rrr d n)"
+| "pseudo_mod_main_list lc r d 0 = r"
+
+
+fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list 
+  \<Rightarrow> nat \<Rightarrow> 'a list \<times> 'a list" where
+  "divmod_poly_one_main_list q r d (Suc n) = (let
+     a = hd r;
+     qqq = cCons a q;
+     rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d))
+     in divmod_poly_one_main_list qqq rr d n)"
+| "divmod_poly_one_main_list q r d 0 = (q,r)"
+
+fun mod_poly_one_main_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list 
+  \<Rightarrow> nat \<Rightarrow> 'a list" where
+  "mod_poly_one_main_list r d (Suc n) = (let
+     a = hd r;
+     rr = tl (if a = 0 then r else minus_poly_rev_list r (map (op * a) d))
+     in mod_poly_one_main_list rr d n)"
+| "mod_poly_one_main_list r d 0 = r"
+
+definition pseudo_divmod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list \<times> 'a list" where
+  "pseudo_divmod_list p q =
+  (if q = [] then ([],p) else
+ (let rq = rev q;
+     (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in 
+   (qu,rev re)))"
+
+definition pseudo_mod_list :: "'a::comm_ring_1 list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+  "pseudo_mod_list p q =
+  (if q = [] then p else
+ (let rq = rev q;
+     re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in 
+   (rev re)))"
+
+lemma minus_zero_does_nothing:
+"(minus_poly_rev_list x (map (op * 0) y)) = (x :: 'a :: ring list)"
+  by(induct x y rule: minus_poly_rev_list.induct, auto)
+
+lemma length_minus_poly_rev_list[simp]:
+ "length (minus_poly_rev_list xs ys) = length xs"
+  by(induct xs ys rule: minus_poly_rev_list.induct, auto)
+
+lemma if_0_minus_poly_rev_list:
+  "(if a = 0 then x else minus_poly_rev_list x (map (op * a) y))
+      = minus_poly_rev_list x (map (op * (a :: 'a :: ring)) y)"
+  by(cases "a=0",simp_all add:minus_zero_does_nothing)
+
+lemma Poly_append:
+  "Poly ((a::'a::comm_semiring_1 list) @ b) = Poly a + monom 1 (length a) * Poly b"
+  by (induct a,auto simp: monom_0 monom_Suc)
+
+lemma minus_poly_rev_list: "length p \<ge> length (q :: 'a :: comm_ring_1 list) \<Longrightarrow>
+  Poly (rev (minus_poly_rev_list (rev p) (rev q)))
+  = Poly p - monom 1 (length p - length q) * Poly q"
+proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct)
+  case (1 x xs y ys) 
+  have "length (rev q) \<le> length (rev p)" using 1 by simp
+  from this[folded 1(2,3)] have ys_xs:"length ys \<le> length xs" by simp
+  hence a:"Poly (rev (minus_poly_rev_list xs ys)) =
+           Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)"
+    by(subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev],auto)
+  have "Poly p - monom 1 (length p - length q) * Poly q
+      = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))"
+    by simp
+  also have "\<dots> = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))"
+    unfolding 1(2,3) by simp
+  also have "\<dots> = Poly (rev xs) + monom x (length xs) -
+  (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" using ys_xs
+    by (simp add:Poly_append distrib_left mult_monom smult_monom)
+  also have "\<dots> = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)"
+    unfolding a diff_monom[symmetric] by(simp)
+  finally show ?case
+    unfolding 1(2,3)[symmetric] by (simp add: smult_monom Poly_append)
+qed auto
+
+lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f"
+  using smult_monom [of a _ n] by (metis mult_smult_left)
+
+lemma head_minus_poly_rev_list:
+  "length d \<le> length r \<Longrightarrow> d\<noteq>[] \<Longrightarrow>
+  hd (minus_poly_rev_list (map (op * (last d :: 'a :: comm_ring)) r) (map (op * (hd r)) (rev d))) = 0"
+proof(induct r)
+  case (Cons a rs)
+  thus ?case by(cases "rev d", simp_all add:ac_simps)
+qed simp
+
+lemma Poly_map: "Poly (map (op * a) p) = smult a (Poly p)"
+proof (induct p)
+  case(Cons x xs) thus ?case by (cases "Poly xs = 0",auto)
+qed simp
+
+lemma last_coeff_is_hd: "xs \<noteq> [] \<Longrightarrow> coeff (Poly xs) (length xs - 1) = hd (rev xs)"
+  by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append)
+
+lemma pseudo_divmod_main_list_invar :
+  assumes leading_nonzero:"last d \<noteq> 0"
+  and lc:"last d = lc"
+  and dNonempty:"d \<noteq> []"
+  and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q',rev r')"
+  and "n = (1 + length r - length d)"
+  shows 
+  "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = 
+  (Poly q', Poly r')"
+using assms(4-)
+proof(induct "n" arbitrary: r q)
+case (Suc n r q)
+  have ifCond: "\<not> Suc (length r) \<le> length d" using Suc.prems by simp
+  have rNonempty:"r \<noteq> []"
+    using ifCond dNonempty using Suc_leI length_greater_0_conv list.size(3) by fastforce
+  let ?a = "(hd (rev r))"
+  let ?rr = "map (op * lc) (rev r)"
+  let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map (op * ?a) (rev d))))"
+  let ?qq = "cCons ?a (map (op * lc) q)"
+  have n: "n = (1 + length r - length d - 1)"
+    using ifCond Suc(3) by simp
+  have rr_val:"(length ?rrr) = (length r - 1)" using ifCond by auto
+  hence rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)"
+    using rNonempty ifCond unfolding One_nat_def by auto
+  from ifCond have id: "Suc (length r) - length d = Suc (length r - length d)" by auto
+  have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')"
+    using Suc.prems ifCond by (simp add:Let_def if_0_minus_poly_rev_list id)
+  hence v:"pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')"
+    using n by auto
+  have sucrr:"Suc (length r) - length d = Suc (length r - length d)"
+    using Suc_diff_le ifCond not_less_eq_eq by blast
+  have n_ok : "n = 1 + (length ?rrr) - length d" using Suc(3) rNonempty by simp
+  have cong: "\<And> x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow>
+    pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp
+  have hd_rev:"coeff (Poly r) (length r - Suc 0) = hd (rev r)"
+    using last_coeff_is_hd[OF rNonempty] by simp
+  show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def
+  proof (rule cong[OF _ _ refl], goal_cases)
+    case 1 
+    show ?case unfolding monom_Suc hd_rev[symmetric]
+      by (simp add: smult_monom Poly_map)
+  next
+    case 2 
+    show ?case 
+    proof (subst Poly_on_rev_starting_with_0, goal_cases)
+      show "hd (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))) = 0"
+        by (fold lc, subst head_minus_poly_rev_list, insert ifCond dNonempty,auto)
+      from ifCond have "length d \<le> length r" by simp
+      then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d =
+        Poly (rev (minus_poly_rev_list (map (op * lc) (rev r)) (map (op * (hd (rev r))) (rev d))))"
+        by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric]
+          minus_poly_rev_list)
+    qed
+  qed simp
+qed simp
+
+lemma pseudo_divmod_impl[code]: "pseudo_divmod (f::'a::comm_ring_1 poly) g =
+  map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))"
+proof (cases "g=0")
+case False
+  hence coeffs_g_nonempty:"(coeffs g) \<noteq> []" by simp
+  hence lastcoeffs:"last (coeffs g) = coeff g (degree g)"
+    by (simp add: hd_rev last_coeffs_eq_coeff_degree not_0_coeffs_not_Nil)
+  obtain q r where qr: "pseudo_divmod_main_list
+            (last (coeffs g)) (rev [])
+            (rev (coeffs f)) (rev (coeffs g))
+            (1 + length (coeffs f) -
+             length (coeffs g)) = (q,rev (rev r))"  by force
+  then have qr': "pseudo_divmod_main_list
+            (hd (rev (coeffs g))) []
+            (rev (coeffs f)) (rev (coeffs g))
+            (1 + length (coeffs f) -
+             length (coeffs g)) = (q,r)" using hd_rev[OF coeffs_g_nonempty] by(auto)
+  from False have cg: "(coeffs g = []) = False" by auto
+  have last_non0:"last (coeffs g) \<noteq> 0" using False by (simp add:last_coeffs_not_0)
+  show ?thesis
+    unfolding pseudo_divmod_def pseudo_divmod_list_def Let_def qr' map_prod_def split cg if_False
+    pseudo_divmod_main_list_invar[OF last_non0 _ _ qr,unfolded lastcoeffs,simplified,symmetric,OF False]
+    poly_of_list_def
+    using False by (auto simp: degree_eq_length_coeffs)
+next
+  case True
+  show ?thesis unfolding True unfolding pseudo_divmod_def pseudo_divmod_list_def
+  by auto
+qed
+
+lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q 
+  xs ys n) = pseudo_mod_main_list l xs ys n"
+  by (induct n arbitrary: l q xs ys, auto simp: Let_def)
+
+lemma pseudo_mod_impl[code]: "pseudo_mod f g =
+  poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))"
+proof -
+  have snd_case: "\<And> f g p. snd ((\<lambda> (x,y). (f x, g y)) p) = g (snd p)" 
+    by auto
+  show ?thesis
+  unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def
+    pseudo_mod_list_def Let_def
+  by (simp add: snd_case pseudo_mod_main_list)
+qed
+
+
+(* *************** *)
+subsubsection \<open>Improved Code-Equations for Polynomial (Pseudo) Division\<close>
+
+lemma pdivmod_via_pseudo_divmod: "pdivmod f g = (if g = 0 then (0,f) 
+     else let 
+       ilc = inverse (coeff g (degree g));       
+       h = smult ilc g;
+       (q,r) = pseudo_divmod f h
+     in (smult ilc q, r))" (is "?l = ?r")
+proof (cases "g = 0")
+  case False
+  define lc where "lc = inverse (coeff g (degree g))"
+  define h where "h = smult lc g"
+  from False have h1: "coeff h (degree h) = 1" and lc: "lc \<noteq> 0" unfolding h_def lc_def by auto
+  hence h0: "h \<noteq> 0" by auto
+  obtain q r where p: "pseudo_divmod f h = (q,r)" by force
+  from False have id: "?r = (smult lc q, r)" 
+    unfolding Let_def h_def[symmetric] lc_def[symmetric] p by auto
+  from pseudo_divmod[OF h0 p, unfolded h1] 
+  have f: "f = h * q + r" and r: "r = 0 \<or> degree r < degree h" by auto
+  have "eucl_rel_poly f h (q, r)" unfolding eucl_rel_poly_iff using f r h0 by auto
+  hence "pdivmod f h = (q,r)" by (simp add: pdivmod_pdivmodrel)
+  hence "pdivmod f g = (smult lc q, r)" 
+    unfolding pdivmod_def h_def div_smult_right[OF lc] mod_smult_right[OF lc]
+    using lc by auto
+  with id show ?thesis by auto
+qed (auto simp: pdivmod_def)
+
+lemma pdivmod_via_pseudo_divmod_list: "pdivmod f g = (let 
+  cg = coeffs g
+  in if cg = [] then (0,f)
+     else let 
+       cf = coeffs f;
+       ilc = inverse (last cg);       
+       ch = map (op * ilc) cg;
+       (q,r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg)
+     in (poly_of_list (map (op * ilc) q), poly_of_list (rev r)))"
+proof -
+  note d = pdivmod_via_pseudo_divmod
+      pseudo_divmod_impl pseudo_divmod_list_def
+  show ?thesis
+  proof (cases "g = 0")
+    case True thus ?thesis unfolding d by auto
+  next
+    case False
+    define ilc where "ilc = inverse (coeff g (degree g))"
+    from False have ilc: "ilc \<noteq> 0" unfolding ilc_def by auto
+    with False have id: "(g = 0) = False" "(coeffs g = []) = False" 
+      "last (coeffs g) = coeff g (degree g)" 
+      "(coeffs (smult ilc g) = []) = False"
+      by (auto simp: last_coeffs_eq_coeff_degree) 
+    have id2: "hd (rev (coeffs (smult ilc g))) = 1"      
+      by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def)
+    have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" 
+      "rev (coeffs (smult ilc g)) = rev (map (op * ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto
+    obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map (op * ilc) (coeffs g)))
+           (1 + length (coeffs f) - length (coeffs g)) = (q,r)" by force
+    show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 
+      unfolding id3 pair map_prod_def split by (auto simp: Poly_map)
+  qed
+qed
+
+lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list"
+proof (intro ext, goal_cases)
+  case (1 q r d n)
+  {
+    fix xs :: "'a list"
+    have "map (op * 1) xs = xs" by (induct xs, auto)
+  } note [simp] = this
+  show ?case by (induct n arbitrary: q r d, auto simp: Let_def)
+qed
+
+fun divide_poly_main_list :: "'a::idom_divide \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list 
+  \<Rightarrow> nat \<Rightarrow> 'a list" where
+  "divide_poly_main_list lc q r d (Suc n) = (let
+     cr = hd r
+     in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let 
+     a = cr div lc;
+     qq = cCons a q;
+     rr = minus_poly_rev_list r (map (op * a) d)
+     in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])"
+| "divide_poly_main_list lc q r d 0 = q"
+
+
+lemma divide_poly_main_list_simp[simp]: "divide_poly_main_list lc q r d (Suc n) = (let
+     cr = hd r;
+     a = cr div lc;
+     qq = cCons a q;
+     rr = minus_poly_rev_list r (map (op * a) d)
+     in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])"
+  by (simp add: Let_def minus_zero_does_nothing)
+
+declare divide_poly_main_list.simps(1)[simp del]
+
+definition divide_poly_list :: "'a::idom_divide poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where
+  "divide_poly_list f g =
+    (let cg = coeffs g
+     in if cg = [] then g
+        else let cf = coeffs f; cgr = rev cg
+          in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))"
+
+lemmas pdivmod_via_divmod_list[code] = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1]
+
+lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n"
+  by  (induct n arbitrary: q r d, auto simp: Let_def)
+
+lemma mod_poly_code[code]: "f mod g =
+    (let cg = coeffs g
+     in if cg = [] then f
+        else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg;
+                 r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg)
+             in poly_of_list (rev r))" (is "?l = ?r")
+proof -
+  have "?l = snd (pdivmod f g)" unfolding pdivmod_def by simp
+  also have "\<dots> = ?r" unfolding pdivmod_via_divmod_list Let_def
+     mod_poly_one_main_list[symmetric, of _ _ _ Nil] by (auto split: prod.splits)
+  finally show ?thesis .
+qed
+
+definition div_field_poly_impl :: "'a :: field poly \<Rightarrow> 'a poly \<Rightarrow> 'a poly" where 
+  "div_field_poly_impl f g = (
+    let cg = coeffs g
+      in if cg = [] then 0
+        else let cf = coeffs f; ilc = inverse (last cg); ch = map (op * ilc) cg;
+                 q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg))
+             in poly_of_list ((map (op * ilc) q)))"
+
+text \<open>We do not declare the following lemma as code equation, since then polynomial division 
+  on non-fields will no longer be executable. However, a code-unfold is possible, since 
+  \<open>div_field_poly_impl\<close> is a bit more efficient than the generic polynomial division.\<close>
+lemma div_field_poly_impl[code_unfold]: "op div = div_field_poly_impl"
+proof (intro ext)
+  fix f g :: "'a poly"
+  have "f div g = fst (pdivmod f g)" unfolding pdivmod_def by simp
+  also have "\<dots> = div_field_poly_impl f g" unfolding 
+    div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits)
+  finally show "f div g =  div_field_poly_impl f g" .
+qed
+
+
+lemma divide_poly_main_list:
+  assumes lc0: "lc \<noteq> 0"
+  and lc:"last d = lc"
+  and d:"d \<noteq> []"
+  and "n = (1 + length r - length d)"
+  shows 
+  "Poly (divide_poly_main_list lc q (rev r) (rev d) n) =
+  divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n"
+using assms(4-)
+proof(induct "n" arbitrary: r q)
+case (Suc n r q)
+  have ifCond: "\<not> Suc (length r) \<le> length d" using Suc.prems by simp
+  have r: "r \<noteq> []"
+    using ifCond d using Suc_leI length_greater_0_conv list.size(3) by fastforce
+  then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases, auto)
+  from d lc obtain dd where d: "d = dd @ [lc]" 
+    by (cases d rule: rev_cases, auto)
+  from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r)
+  from ifCond have len: "length dd \<le> length rr" by (simp add: r d)
+  show ?case
+  proof (cases "lcr div lc * lc = lcr")
+    case False
+    thus ?thesis unfolding Suc(2)[symmetric] using r d
+      by (auto simp add: Let_def nth_default_append)
+  next
+    case True
+    hence id:
+    "?thesis = (Poly (divide_poly_main_list lc (cCons (lcr div lc) q)
+         (rev (rev (minus_poly_rev_list (rev rr) (rev (map (op * (lcr div lc)) dd))))) (rev d) n) = 
+      divide_poly_main lc
+           (monom 1 (Suc n) * Poly q + monom (lcr div lc) n)
+           (Poly r - monom (lcr div lc) n * Poly d)
+           (Poly d) (length rr - 1) n)"
+           using r d 
+      by (cases r rule: rev_cases; cases "d" rule: rev_cases; 
+        auto simp add: Let_def rev_map nth_default_append)      
+    have cong: "\<And> x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \<Longrightarrow> x2 = y2 \<Longrightarrow> x3 = y3 \<Longrightarrow> x4 = y4 \<Longrightarrow>
+      divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp
+    show ?thesis unfolding id 
+    proof (subst Suc(1), simp add: n,
+      subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases)
+      case 2 
+      have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)"
+        by (simp add: mult_monom len True)
+      thus ?case unfolding r d Poly_append n ring_distribs
+        by (auto simp: Poly_map smult_monom smult_monom_mult)
+    qed (auto simp: len monom_Suc smult_monom)
+  qed
+qed simp
+
+
+lemma divide_poly_list[code]: "f div g = divide_poly_list f g" 
+proof -
+  note d = divide_poly_def divide_poly_list_def
+  show ?thesis
+  proof (cases "g = 0")
+    case True
+    show ?thesis unfolding d True by auto
+  next
+    case False
+    then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases, auto)    
+    with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto
+    from cg False have lcg: "coeff g (degree g) = lcg" 
+      using last_coeffs_eq_coeff_degree last_snoc by force
+    with False have lcg0: "lcg \<noteq> 0" by auto
+    from cg have ltp: "Poly (cg @ [lcg]) = g"
+     using Poly_coeffs [of g] by auto
+    show ?thesis unfolding d cg Let_def id if_False poly_of_list_def
+      by (subst divide_poly_main_list, insert False cg lcg0, auto simp: lcg ltp,
+      simp add: degree_eq_length_coeffs)
+  qed
 qed
 
 no_notation cCons (infixr "##" 65)
--- a/src/HOL/Library/Polynomial_Factorial.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Library/Polynomial_Factorial.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -19,40 +19,6 @@
 lemma prod_mset_const_poly: "prod_mset (image_mset (\<lambda>x. [:f x:]) A) = [:prod_mset (image_mset f A):]"
   by (induction A) (simp_all add: one_poly_def mult_ac)
 
-lemma is_unit_smult_iff: "smult c p dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
-proof -
-  have "smult c p = [:c:] * p" by simp
-  also have "\<dots> dvd 1 \<longleftrightarrow> c dvd 1 \<and> p dvd 1"
-  proof safe
-    assume A: "[:c:] * p dvd 1"
-    thus "p dvd 1" by (rule dvd_mult_right)
-    from A obtain q where B: "1 = [:c:] * p * q" by (erule dvdE)
-    have "c dvd c * (coeff p 0 * coeff q 0)" by simp
-    also have "\<dots> = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult)
-    also note B [symmetric]
-    finally show "c dvd 1" by simp
-  next
-    assume "c dvd 1" "p dvd 1"
-    from \<open>c dvd 1\<close> obtain d where "1 = c * d" by (erule dvdE)
-    hence "1 = [:c:] * [:d:]" by (simp add: one_poly_def mult_ac)
-    hence "[:c:] dvd 1" by (rule dvdI)
-    from mult_dvd_mono[OF this \<open>p dvd 1\<close>] show "[:c:] * p dvd 1" by simp
-  qed
-  finally show ?thesis .
-qed
-
-lemma degree_mod_less': "b \<noteq> 0 \<Longrightarrow> a mod b \<noteq> 0 \<Longrightarrow> degree (a mod b) < degree b"
-  using degree_mod_less[of b a] by auto
-  
-lemma smult_eq_iff:
-  assumes "(b :: 'a :: field) \<noteq> 0"
-  shows   "smult a p = smult b q \<longleftrightarrow> smult (a / b) p = q"
-proof
-  assume "smult a p = smult b q"
-  also from assms have "smult (inverse b) \<dots> = q" by simp
-  finally show "smult (a / b) p = q" by (simp add: field_simps)
-qed (insert assms, auto)
-
 lemma irreducible_const_poly_iff:
   fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}"
   shows "irreducible [:c:] \<longleftrightarrow> irreducible c"
@@ -160,145 +126,6 @@
   by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract)
 
 
-subsection \<open>Content and primitive part of a polynomial\<close>
-
-definition content :: "('a :: semiring_Gcd poly) \<Rightarrow> 'a" where
-  "content p = Gcd (set (coeffs p))"
-
-lemma content_0 [simp]: "content 0 = 0"
-  by (simp add: content_def)
-
-lemma content_1 [simp]: "content 1 = 1"
-  by (simp add: content_def)
-
-lemma content_const [simp]: "content [:c:] = normalize c"
-  by (simp add: content_def cCons_def)
-
-lemma const_poly_dvd_iff_dvd_content:
-  fixes c :: "'a :: semiring_Gcd"
-  shows "[:c:] dvd p \<longleftrightarrow> c dvd content p"
-proof (cases "p = 0")
-  case [simp]: False
-  have "[:c:] dvd p \<longleftrightarrow> (\<forall>n. c dvd coeff p n)" by (rule const_poly_dvd_iff)
-  also have "\<dots> \<longleftrightarrow> (\<forall>a\<in>set (coeffs p). c dvd a)"
-  proof safe
-    fix n :: nat assume "\<forall>a\<in>set (coeffs p). c dvd a"
-    thus "c dvd coeff p n"
-      by (cases "n \<le> degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits)
-  qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits)
-  also have "\<dots> \<longleftrightarrow> c dvd content p"
-    by (simp add: content_def dvd_Gcd_iff mult.commute [of "unit_factor x" for x]
-          dvd_mult_unit_iff lead_coeff_nonzero)
-  finally show ?thesis .
-qed simp_all
-
-lemma content_dvd [simp]: "[:content p:] dvd p"
-  by (subst const_poly_dvd_iff_dvd_content) simp_all
-  
-lemma content_dvd_coeff [simp]: "content p dvd coeff p n"
-  by (cases "n \<le> degree p") 
-     (auto simp: content_def coeffs_def not_le coeff_eq_0 simp del: upt_Suc intro: Gcd_dvd)
-
-lemma content_dvd_coeffs: "c \<in> set (coeffs p) \<Longrightarrow> content p dvd c"
-  by (simp add: content_def Gcd_dvd)
-
-lemma normalize_content [simp]: "normalize (content p) = content p"
-  by (simp add: content_def)
-
-lemma is_unit_content_iff [simp]: "is_unit (content p) \<longleftrightarrow> content p = 1"
-proof
-  assume "is_unit (content p)"
-  hence "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content)
-  thus "content p = 1" by simp
-qed auto
-
-lemma content_smult [simp]: "content (smult c p) = normalize c * content p"
-  by (simp add: content_def coeffs_smult Gcd_mult)
-
-lemma content_eq_zero_iff [simp]: "content p = 0 \<longleftrightarrow> p = 0"
-  by (auto simp: content_def simp: poly_eq_iff coeffs_def)
-
-definition primitive_part :: "'a :: {semiring_Gcd,idom_divide} poly \<Rightarrow> 'a poly" where
-  "primitive_part p = (if p = 0 then 0 else map_poly (\<lambda>x. x div content p) p)"
-  
-lemma primitive_part_0 [simp]: "primitive_part 0 = 0"
-  by (simp add: primitive_part_def)
-
-lemma content_times_primitive_part [simp]:
-  fixes p :: "'a :: {idom_divide, semiring_Gcd} poly"
-  shows "smult (content p) (primitive_part p) = p"
-proof (cases "p = 0")
-  case False
-  thus ?thesis
-  unfolding primitive_part_def
-  by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs 
-           intro: map_poly_idI)
-qed simp_all
-
-lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \<longleftrightarrow> p = 0"
-proof (cases "p = 0")
-  case False
-  hence "primitive_part p = map_poly (\<lambda>x. x div content p) p"
-    by (simp add:  primitive_part_def)
-  also from False have "\<dots> = 0 \<longleftrightarrow> p = 0"
-    by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs)
-  finally show ?thesis using False by simp
-qed simp
-
-lemma content_primitive_part [simp]:
-  assumes "p \<noteq> 0"
-  shows   "content (primitive_part p) = 1"
-proof -
-  have "p = smult (content p) (primitive_part p)" by simp
-  also have "content \<dots> = content p * content (primitive_part p)" 
-    by (simp del: content_times_primitive_part)
-  finally show ?thesis using assms by simp
-qed
-
-lemma content_decompose:
-  fixes p :: "'a :: semiring_Gcd poly"
-  obtains p' where "p = smult (content p) p'" "content p' = 1"
-proof (cases "p = 0")
-  case True
-  thus ?thesis by (intro that[of 1]) simp_all
-next
-  case False
-  from content_dvd[of p] obtain r where r: "p = [:content p:] * r" by (erule dvdE)
-  have "content p * 1 = content p * content r" by (subst r) simp
-  with False have "content r = 1" by (subst (asm) mult_left_cancel) simp_all
-  with r show ?thesis by (intro that[of r]) simp_all
-qed
-
-lemma smult_content_normalize_primitive_part [simp]:
-  "smult (content p) (normalize (primitive_part p)) = normalize p"
-proof -
-  have "smult (content p) (normalize (primitive_part p)) = 
-          normalize ([:content p:] * primitive_part p)" 
-    by (subst normalize_mult) (simp_all add: normalize_const_poly)
-  also have "[:content p:] * primitive_part p = p" by simp
-  finally show ?thesis .
-qed
-
-lemma content_dvd_contentI [intro]:
-  "p dvd q \<Longrightarrow> content p dvd content q"
-  using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast
-  
-lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]"
-  by (simp add: primitive_part_def map_poly_pCons)
- 
-lemma primitive_part_prim: "content p = 1 \<Longrightarrow> primitive_part p = p"
-  by (auto simp: primitive_part_def)
-  
-lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p"
-proof (cases "p = 0")
-  case False
-  have "p = smult (content p) (primitive_part p)" by simp
-  also from False have "degree \<dots> = degree (primitive_part p)"
-    by (subst degree_smult_eq) simp_all
-  finally show ?thesis ..
-qed simp_all
-
-
 subsection \<open>Lifting polynomial coefficients to the field of fractions\<close>
 
 abbreviation (input) fract_poly 
@@ -690,7 +517,7 @@
   fix p :: "'a poly"
   show "unit_factor_field_poly p * normalize_field_poly p = p"
     by (cases "p = 0") 
-       (simp_all add: unit_factor_field_poly_def normalize_field_poly_def lead_coeff_nonzero)
+       (simp_all add: unit_factor_field_poly_def normalize_field_poly_def)
 next
   fix p :: "'a poly" assume "is_unit p"
   thus "normalize_field_poly p = 1"
@@ -698,7 +525,7 @@
 next
   fix p :: "'a poly" assume "p \<noteq> 0"
   thus "is_unit (unit_factor_field_poly p)"
-    by (simp add: unit_factor_field_poly_def lead_coeff_nonzero is_unit_pCons_iff)
+    by (simp add: unit_factor_field_poly_def is_unit_pCons_iff)
 next
   fix p q s :: "'a poly" assume "s \<noteq> 0"
   moreover assume "euclidean_size_field_poly p < euclidean_size_field_poly q"
@@ -943,7 +770,7 @@
   have "fract_poly p = unit_factor_field_poly (fract_poly p) * 
           normalize_field_poly (fract_poly p)" by simp
   also have "unit_factor_field_poly (fract_poly p) = [:to_fract (lead_coeff p):]" 
-    by (simp add: unit_factor_field_poly_def lead_coeff_def monom_0 degree_map_poly coeff_map_poly)
+    by (simp add: unit_factor_field_poly_def monom_0 degree_map_poly coeff_map_poly)
   also from assms have "normalize_field_poly (fract_poly p) = prod_mset ?P" 
     by (subst field_poly_prod_mset_prime_factorization) simp_all
   also have "\<dots> = prod_mset (image_mset id ?P)" by simp
--- a/src/HOL/Real_Vector_Spaces.thy	Thu Jan 05 22:37:52 2017 +0100
+++ b/src/HOL/Real_Vector_Spaces.thy	Thu Jan 05 22:38:06 2017 +0100
@@ -176,6 +176,21 @@
   for x :: "'a::real_vector"
   using scaleR_minus_left [of 1 x] by simp
 
+lemma scaleR_2:
+  fixes x :: "'a::real_vector"
+  shows "scaleR 2 x = x + x"
+  unfolding one_add_one [symmetric] scaleR_left_distrib by simp
+
+lemma scaleR_half_double [simp]:
+  fixes a :: "'a::real_vector"
+  shows "(1 / 2) *\<^sub>R (a + a) = a"
+proof -
+  have "\<And>r. r *\<^sub>R (a + a) = (r * 2) *\<^sub>R a"
+    by (metis scaleR_2 scaleR_scaleR)
+  then show ?thesis
+    by simp
+qed
+
 class real_algebra = real_vector + ring +
   assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
     and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"